| author | wenzelm | 
| Sat, 13 Oct 2001 20:32:07 +0200 | |
| changeset 11741 | 470e608d7a74 | 
| parent 11655 | 923e4d0d36d5 | 
| child 11753 | 02b257ef0ee2 | 
| permissions | -rw-r--r-- | 
| 9422 | 1  | 
(* Title: HOL/Set.ML  | 
| 923 | 2  | 
ID: $Id$  | 
| 1465 | 3  | 
Author: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
| 923 | 4  | 
Copyright 1991 University of Cambridge  | 
5  | 
||
| 
1985
 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 
paulson 
parents: 
1937 
diff
changeset
 | 
6  | 
Set theory for higher-order logic. A set is simply a predicate.  | 
| 923 | 7  | 
*)  | 
8  | 
||
| 1548 | 9  | 
section "Relating predicates and sets";  | 
10  | 
||
| 
3469
 
61d927bd57ec
Now Collect_mem_eq is a default simprule (how could it have ever been omitted?
 
paulson 
parents: 
3420 
diff
changeset
 | 
11  | 
Addsimps [Collect_mem_eq];  | 
| 
 
61d927bd57ec
Now Collect_mem_eq is a default simprule (how could it have ever been omitted?
 
paulson 
parents: 
3420 
diff
changeset
 | 
12  | 
AddIffs [mem_Collect_eq];  | 
| 
2499
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
13  | 
|
| 
5143
 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 
paulson 
parents: 
5069 
diff
changeset
 | 
14  | 
Goal "P(a) ==> a : {x. P(x)}";
 | 
| 
2499
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
15  | 
by (Asm_simp_tac 1);  | 
| 923 | 16  | 
qed "CollectI";  | 
17  | 
||
| 5316 | 18  | 
Goal "a : {x. P(x)} ==> P(a)";
 | 
| 
2499
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
19  | 
by (Asm_full_simp_tac 1);  | 
| 923 | 20  | 
qed "CollectD";  | 
21  | 
||
| 
9687
 
772ac061bd76
moved congruence rules UN_cong, INT_cong from UNTIY/Union to Set.ML
 
paulson 
parents: 
9422 
diff
changeset
 | 
22  | 
val [prem] = Goal "(!!x. (x:A) = (x:B)) ==> A = B";  | 
| 923 | 23  | 
by (rtac (prem RS ext RS arg_cong RS box_equals) 1);  | 
24  | 
by (rtac Collect_mem_eq 1);  | 
|
25  | 
by (rtac Collect_mem_eq 1);  | 
|
26  | 
qed "set_ext";  | 
|
27  | 
||
| 
9687
 
772ac061bd76
moved congruence rules UN_cong, INT_cong from UNTIY/Union to Set.ML
 
paulson 
parents: 
9422 
diff
changeset
 | 
28  | 
val [prem] = Goal "(!!x. P(x)=Q(x)) ==> {x. P(x)} = {x. Q(x)}";
 | 
| 923 | 29  | 
by (rtac (prem RS ext RS arg_cong) 1);  | 
30  | 
qed "Collect_cong";  | 
|
31  | 
||
| 9108 | 32  | 
bind_thm ("CollectE", make_elim CollectD);
 | 
| 923 | 33  | 
|
| 
2499
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
34  | 
AddSIs [CollectI];  | 
| 
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
35  | 
AddSEs [CollectE];  | 
| 
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
36  | 
|
| 
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
37  | 
|
| 1548 | 38  | 
section "Bounded quantifiers";  | 
| 923 | 39  | 
|
| 5316 | 40  | 
val prems = Goalw [Ball_def]  | 
| 9041 | 41  | 
"[| !!x. x:A ==> P(x) |] ==> ALL x:A. P(x)";  | 
| 923 | 42  | 
by (REPEAT (ares_tac (prems @ [allI,impI]) 1));  | 
43  | 
qed "ballI";  | 
|
44  | 
||
| 8839 | 45  | 
bind_thms ("strip", [impI, allI, ballI]);
 | 
46  | 
||
| 9041 | 47  | 
Goalw [Ball_def] "[| ALL x:A. P(x); x:A |] ==> P(x)";  | 
| 5316 | 48  | 
by (Blast_tac 1);  | 
| 923 | 49  | 
qed "bspec";  | 
50  | 
||
| 5316 | 51  | 
val major::prems = Goalw [Ball_def]  | 
| 9041 | 52  | 
"[| ALL x:A. P(x); P(x) ==> Q; x~:A ==> Q |] ==> Q";  | 
| 923 | 53  | 
by (rtac (major RS spec RS impCE) 1);  | 
54  | 
by (REPEAT (eresolve_tac prems 1));  | 
|
55  | 
qed "ballE";  | 
|
56  | 
||
| 9041 | 57  | 
(*Takes assumptions ALL x:A.P(x) and a:A; creates assumption P(a)*)  | 
| 923 | 58  | 
fun ball_tac i = etac ballE i THEN contr_tac (i+1);  | 
59  | 
||
| 
2499
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
60  | 
AddSIs [ballI];  | 
| 
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
61  | 
AddEs [ballE];  | 
| 7441 | 62  | 
AddXDs [bspec];  | 
| 5521 | 63  | 
(* gives better instantiation for bound: *)  | 
| 11166 | 64  | 
claset_ref() := claset() addbefore ("bspec", datac bspec 1);
 | 
| 
2499
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
65  | 
|
| 6006 | 66  | 
(*Normally the best argument order: P(x) constrains the choice of x:A*)  | 
| 9041 | 67  | 
Goalw [Bex_def] "[| P(x); x:A |] ==> EX x:A. P(x)";  | 
| 5316 | 68  | 
by (Blast_tac 1);  | 
| 923 | 69  | 
qed "bexI";  | 
70  | 
||
| 6006 | 71  | 
(*The best argument order when there is only one x:A*)  | 
| 9041 | 72  | 
Goalw [Bex_def] "[| x:A; P(x) |] ==> EX x:A. P(x)";  | 
| 6006 | 73  | 
by (Blast_tac 1);  | 
74  | 
qed "rev_bexI";  | 
|
75  | 
||
| 7031 | 76  | 
val prems = Goal  | 
| 9041 | 77  | 
"[| ALL x:A. ~P(x) ==> P(a); a:A |] ==> EX x:A. P(x)";  | 
| 7007 | 78  | 
by (rtac classical 1);  | 
79  | 
by (REPEAT (ares_tac (prems@[bexI,ballI,notI,notE]) 1)) ;  | 
|
80  | 
qed "bexCI";  | 
|
| 923 | 81  | 
|
| 5316 | 82  | 
val major::prems = Goalw [Bex_def]  | 
| 9041 | 83  | 
"[| EX x:A. P(x); !!x. [| x:A; P(x) |] ==> Q |] ==> Q";  | 
| 923 | 84  | 
by (rtac (major RS exE) 1);  | 
85  | 
by (REPEAT (eresolve_tac (prems @ [asm_rl,conjE]) 1));  | 
|
86  | 
qed "bexE";  | 
|
87  | 
||
| 
2499
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
88  | 
AddIs [bexI];  | 
| 
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
89  | 
AddSEs [bexE];  | 
| 
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
90  | 
|
| 3420 | 91  | 
(*Trival rewrite rule*)  | 
| 9041 | 92  | 
Goal "(ALL x:A. P) = ((EX x. x:A) --> P)";  | 
| 4089 | 93  | 
by (simp_tac (simpset() addsimps [Ball_def]) 1);  | 
| 3420 | 94  | 
qed "ball_triv";  | 
| 1816 | 95  | 
|
| 1882 | 96  | 
(*Dual form for existentials*)  | 
| 9041 | 97  | 
Goal "(EX x:A. P) = ((EX x. x:A) & P)";  | 
| 4089 | 98  | 
by (simp_tac (simpset() addsimps [Bex_def]) 1);  | 
| 3420 | 99  | 
qed "bex_triv";  | 
| 1882 | 100  | 
|
| 3420 | 101  | 
Addsimps [ball_triv, bex_triv];  | 
| 923 | 102  | 
|
| 
11223
 
fef9da0ee890
I forgot a few bases cases for the 1-point rules...
 
nipkow 
parents: 
11220 
diff
changeset
 | 
103  | 
Goal "(EX x:A. x=a) = (a:A)";  | 
| 11220 | 104  | 
by(Blast_tac 1);  | 
| 
11223
 
fef9da0ee890
I forgot a few bases cases for the 1-point rules...
 
nipkow 
parents: 
11220 
diff
changeset
 | 
105  | 
qed "bex_triv_one_point1";  | 
| 
 
fef9da0ee890
I forgot a few bases cases for the 1-point rules...
 
nipkow 
parents: 
11220 
diff
changeset
 | 
106  | 
|
| 
 
fef9da0ee890
I forgot a few bases cases for the 1-point rules...
 
nipkow 
parents: 
11220 
diff
changeset
 | 
107  | 
Goal "(EX x:A. a=x) = (a:A)";  | 
| 
 
fef9da0ee890
I forgot a few bases cases for the 1-point rules...
 
nipkow 
parents: 
11220 
diff
changeset
 | 
108  | 
by(Blast_tac 1);  | 
| 
 
fef9da0ee890
I forgot a few bases cases for the 1-point rules...
 
nipkow 
parents: 
11220 
diff
changeset
 | 
109  | 
qed "bex_triv_one_point2";  | 
| 11220 | 110  | 
|
111  | 
Goal "(EX x:A. x=a & P x) = (a:A & P a)";  | 
|
112  | 
by(Blast_tac 1);  | 
|
| 
11223
 
fef9da0ee890
I forgot a few bases cases for the 1-point rules...
 
nipkow 
parents: 
11220 
diff
changeset
 | 
113  | 
qed "bex_one_point1";  | 
| 
 
fef9da0ee890
I forgot a few bases cases for the 1-point rules...
 
nipkow 
parents: 
11220 
diff
changeset
 | 
114  | 
|
| 11232 | 115  | 
Goal "(EX x:A. a=x & P x) = (a:A & P a)";  | 
116  | 
by(Blast_tac 1);  | 
|
117  | 
qed "bex_one_point2";  | 
|
118  | 
||
| 
11223
 
fef9da0ee890
I forgot a few bases cases for the 1-point rules...
 
nipkow 
parents: 
11220 
diff
changeset
 | 
119  | 
Goal "(ALL x:A. x=a --> P x) = (a:A --> P a)";  | 
| 
 
fef9da0ee890
I forgot a few bases cases for the 1-point rules...
 
nipkow 
parents: 
11220 
diff
changeset
 | 
120  | 
by(Blast_tac 1);  | 
| 
 
fef9da0ee890
I forgot a few bases cases for the 1-point rules...
 
nipkow 
parents: 
11220 
diff
changeset
 | 
121  | 
qed "ball_one_point1";  | 
| 11220 | 122  | 
|
| 
11223
 
fef9da0ee890
I forgot a few bases cases for the 1-point rules...
 
nipkow 
parents: 
11220 
diff
changeset
 | 
123  | 
Goal "(ALL x:A. a=x --> P x) = (a:A --> P a)";  | 
| 
 
fef9da0ee890
I forgot a few bases cases for the 1-point rules...
 
nipkow 
parents: 
11220 
diff
changeset
 | 
124  | 
by(Blast_tac 1);  | 
| 
 
fef9da0ee890
I forgot a few bases cases for the 1-point rules...
 
nipkow 
parents: 
11220 
diff
changeset
 | 
125  | 
qed "ball_one_point2";  | 
| 
 
fef9da0ee890
I forgot a few bases cases for the 1-point rules...
 
nipkow 
parents: 
11220 
diff
changeset
 | 
126  | 
|
| 11232 | 127  | 
Addsimps [bex_triv_one_point1,bex_triv_one_point2,  | 
128  | 
bex_one_point1,bex_one_point2,  | 
|
| 
11223
 
fef9da0ee890
I forgot a few bases cases for the 1-point rules...
 
nipkow 
parents: 
11220 
diff
changeset
 | 
129  | 
ball_one_point1,ball_one_point2];  | 
| 11220 | 130  | 
|
131  | 
let  | 
|
132  | 
val ex_pattern = Thm.read_cterm (Theory.sign_of (the_context ()))  | 
|
133  | 
    ("EX x:A. P(x) & Q(x)",HOLogic.boolT)
 | 
|
134  | 
||
135  | 
val prove_bex_tac = rewrite_goals_tac [Bex_def] THEN  | 
|
136  | 
Quantifier1.prove_one_point_ex_tac;  | 
|
137  | 
||
138  | 
val rearrange_bex = Quantifier1.rearrange_bex prove_bex_tac;  | 
|
139  | 
||
140  | 
val all_pattern = Thm.read_cterm (Theory.sign_of (the_context ()))  | 
|
| 11232 | 141  | 
    ("ALL x:A. P(x) --> Q(x)",HOLogic.boolT)
 | 
| 11220 | 142  | 
|
| 11232 | 143  | 
val prove_ball_tac = rewrite_goals_tac [Ball_def] THEN  | 
| 11220 | 144  | 
Quantifier1.prove_one_point_all_tac;  | 
145  | 
||
146  | 
val rearrange_ball = Quantifier1.rearrange_ball prove_ball_tac;  | 
|
147  | 
||
148  | 
val defBEX_regroup = mk_simproc "defined BEX" [ex_pattern] rearrange_bex;  | 
|
149  | 
val defBALL_regroup = mk_simproc "defined BALL" [all_pattern] rearrange_ball;  | 
|
150  | 
in  | 
|
151  | 
||
152  | 
Addsimprocs [defBALL_regroup,defBEX_regroup]  | 
|
153  | 
||
154  | 
end;  | 
|
155  | 
||
| 923 | 156  | 
(** Congruence rules **)  | 
157  | 
||
| 6291 | 158  | 
val prems = Goalw [Ball_def]  | 
| 923 | 159  | 
"[| A=B; !!x. x:B ==> P(x) = Q(x) |] ==> \  | 
| 9041 | 160  | 
\ (ALL x:A. P(x)) = (ALL x:B. Q(x))";  | 
| 6291 | 161  | 
by (asm_simp_tac (simpset() addsimps prems) 1);  | 
| 923 | 162  | 
qed "ball_cong";  | 
163  | 
||
| 6291 | 164  | 
val prems = Goalw [Bex_def]  | 
| 923 | 165  | 
"[| A=B; !!x. x:B ==> P(x) = Q(x) |] ==> \  | 
| 9041 | 166  | 
\ (EX x:A. P(x)) = (EX x:B. Q(x))";  | 
| 6291 | 167  | 
by (asm_simp_tac (simpset() addcongs [conj_cong] addsimps prems) 1);  | 
| 923 | 168  | 
qed "bex_cong";  | 
169  | 
||
| 6291 | 170  | 
Addcongs [ball_cong,bex_cong];  | 
171  | 
||
| 1548 | 172  | 
section "Subsets";  | 
| 923 | 173  | 
|
| 5316 | 174  | 
val prems = Goalw [subset_def] "(!!x. x:A ==> x:B) ==> A <= B";  | 
| 923 | 175  | 
by (REPEAT (ares_tac (prems @ [ballI]) 1));  | 
176  | 
qed "subsetI";  | 
|
177  | 
||
| 5649 | 178  | 
(*Map the type ('a set => anything) to just 'a.
 | 
179  | 
For overloading constants whose first argument has type "'a set" *)  | 
|
180  | 
fun overload_1st_set s = Blast.overloaded (s, HOLogic.dest_setT o domain_type);  | 
|
181  | 
||
| 4059 | 182  | 
(*While (:) is not, its type must be kept  | 
183  | 
for overloading of = to work.*)  | 
|
| 
4240
 
8ba60a4cd380
Renamed "overload" to "overloaded" for sml/nj compatibility
 
paulson 
parents: 
4231 
diff
changeset
 | 
184  | 
Blast.overloaded ("op :", domain_type);
 | 
| 5649 | 185  | 
|
186  | 
overload_1st_set "Ball"; (*need UNION, INTER also?*)  | 
|
187  | 
overload_1st_set "Bex";  | 
|
| 4059 | 188  | 
|
| 4469 | 189  | 
(*Image: retain the type of the set being expressed*)  | 
| 8005 | 190  | 
Blast.overloaded ("image", domain_type);
 | 
| 2881 | 191  | 
|
| 923 | 192  | 
(*Rule in Modus Ponens style*)  | 
| 5316 | 193  | 
Goalw [subset_def] "[| A <= B; c:A |] ==> c:B";  | 
194  | 
by (Blast_tac 1);  | 
|
| 923 | 195  | 
qed "subsetD";  | 
| 7658 | 196  | 
AddXIs [subsetD];  | 
| 923 | 197  | 
|
198  | 
(*The same, with reversed premises for use with etac -- cf rev_mp*)  | 
|
| 7007 | 199  | 
Goal "[| c:A; A <= B |] ==> c:B";  | 
200  | 
by (REPEAT (ares_tac [subsetD] 1)) ;  | 
|
201  | 
qed "rev_subsetD";  | 
|
| 7658 | 202  | 
AddXIs [rev_subsetD];  | 
| 923 | 203  | 
|
| 1920 | 204  | 
(*Converts A<=B to x:A ==> x:B*)  | 
205  | 
fun impOfSubs th = th RSN (2, rev_subsetD);  | 
|
206  | 
||
| 923 | 207  | 
(*Classical elimination rule*)  | 
| 5316 | 208  | 
val major::prems = Goalw [subset_def]  | 
| 923 | 209  | 
"[| A <= B; c~:A ==> P; c:B ==> P |] ==> P";  | 
210  | 
by (rtac (major RS ballE) 1);  | 
|
211  | 
by (REPEAT (eresolve_tac prems 1));  | 
|
212  | 
qed "subsetCE";  | 
|
213  | 
||
214  | 
(*Takes assumptions A<=B; c:A and creates the assumption c:B *)  | 
|
215  | 
fun set_mp_tac i = etac subsetCE i THEN mp_tac i;  | 
|
216  | 
||
| 
2499
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
217  | 
AddSIs [subsetI];  | 
| 
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
218  | 
AddEs [subsetD, subsetCE];  | 
| 923 | 219  | 
|
| 10233 | 220  | 
Goal "[| A <= B; c ~: B |] ==> c ~: A";  | 
221  | 
by (Blast_tac 1);  | 
|
222  | 
qed "contra_subsetD";  | 
|
223  | 
||
| 7007 | 224  | 
Goal "A <= (A::'a set)";  | 
225  | 
by (Fast_tac 1);  | 
|
| 10233 | 226  | 
qed "subset_refl";  | 
| 
2499
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
227  | 
|
| 5316 | 228  | 
Goal "[| A<=B; B<=C |] ==> A<=(C::'a set)";  | 
| 2891 | 229  | 
by (Blast_tac 1);  | 
| 923 | 230  | 
qed "subset_trans";  | 
231  | 
||
232  | 
||
| 1548 | 233  | 
section "Equality";  | 
| 923 | 234  | 
|
235  | 
(*Anti-symmetry of the subset relation*)  | 
|
| 5316 | 236  | 
Goal "[| A <= B; B <= A |] ==> A = (B::'a set)";  | 
| 5318 | 237  | 
by (rtac set_ext 1);  | 
| 5316 | 238  | 
by (blast_tac (claset() addIs [subsetD]) 1);  | 
| 923 | 239  | 
qed "subset_antisym";  | 
| 9108 | 240  | 
bind_thm ("equalityI", subset_antisym);
 | 
| 923 | 241  | 
|
| 1762 | 242  | 
AddSIs [equalityI];  | 
243  | 
||
| 923 | 244  | 
(* Equality rules from ZF set theory -- are they appropriate here? *)  | 
| 5316 | 245  | 
Goal "A = B ==> A<=(B::'a set)";  | 
246  | 
by (etac ssubst 1);  | 
|
| 923 | 247  | 
by (rtac subset_refl 1);  | 
248  | 
qed "equalityD1";  | 
|
249  | 
||
| 5316 | 250  | 
Goal "A = B ==> B<=(A::'a set)";  | 
251  | 
by (etac ssubst 1);  | 
|
| 923 | 252  | 
by (rtac subset_refl 1);  | 
253  | 
qed "equalityD2";  | 
|
254  | 
||
| 9338 | 255  | 
(*Be careful when adding this to the claset as subset_empty is in the simpset:  | 
256  | 
  A={} goes to {}<=A and A<={} and then back to A={} !*)
 | 
|
| 5316 | 257  | 
val prems = Goal  | 
| 923 | 258  | 
"[| A = B; [| A<=B; B<=(A::'a set) |] ==> P |] ==> P";  | 
259  | 
by (resolve_tac prems 1);  | 
|
260  | 
by (REPEAT (resolve_tac (prems RL [equalityD1,equalityD2]) 1));  | 
|
261  | 
qed "equalityE";  | 
|
262  | 
||
| 5316 | 263  | 
val major::prems = Goal  | 
| 923 | 264  | 
"[| A = B; [| c:A; c:B |] ==> P; [| c~:A; c~:B |] ==> P |] ==> P";  | 
265  | 
by (rtac (major RS equalityE) 1);  | 
|
266  | 
by (REPEAT (contr_tac 1 ORELSE eresolve_tac ([asm_rl,subsetCE]@prems) 1));  | 
|
267  | 
qed "equalityCE";  | 
|
268  | 
||
| 
9186
 
7b2f4e6538b4
now uses equalityCE, which usually is more efficent than equalityE
 
paulson 
parents: 
9161 
diff
changeset
 | 
269  | 
AddEs [equalityCE];  | 
| 
 
7b2f4e6538b4
now uses equalityCE, which usually is more efficent than equalityE
 
paulson 
parents: 
9161 
diff
changeset
 | 
270  | 
|
| 923 | 271  | 
(*Lemma for creating induction formulae -- for "pattern matching" on p  | 
272  | 
To make the induction hypotheses usable, apply "spec" or "bspec" to  | 
|
273  | 
put universal quantifiers over the free variables in p. *)  | 
|
| 5316 | 274  | 
val prems = Goal  | 
| 923 | 275  | 
"[| p:A; !!z. z:A ==> p=z --> R |] ==> R";  | 
276  | 
by (rtac mp 1);  | 
|
277  | 
by (REPEAT (resolve_tac (refl::prems) 1));  | 
|
278  | 
qed "setup_induction";  | 
|
279  | 
||
| 8053 | 280  | 
Goal "A = B ==> (x : A) = (x : B)";  | 
281  | 
by (Asm_simp_tac 1);  | 
|
282  | 
qed "eqset_imp_iff";  | 
|
283  | 
||
| 923 | 284  | 
|
| 
4159
 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 
paulson 
parents: 
4135 
diff
changeset
 | 
285  | 
section "The universal set -- UNIV";  | 
| 
 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 
paulson 
parents: 
4135 
diff
changeset
 | 
286  | 
|
| 7031 | 287  | 
Goalw [UNIV_def] "x : UNIV";  | 
288  | 
by (rtac CollectI 1);  | 
|
289  | 
by (rtac TrueI 1);  | 
|
290  | 
qed "UNIV_I";  | 
|
| 
4159
 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 
paulson 
parents: 
4135 
diff
changeset
 | 
291  | 
|
| 4434 | 292  | 
Addsimps [UNIV_I];  | 
293  | 
AddIs [UNIV_I]; (*unsafe makes it less likely to cause problems*)  | 
|
| 
4159
 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 
paulson 
parents: 
4135 
diff
changeset
 | 
294  | 
|
| 10482 | 295  | 
Goal "EX x. x : UNIV";  | 
296  | 
by (Simp_tac 1);  | 
|
297  | 
qed "UNIV_witness";  | 
|
298  | 
AddXIs [UNIV_witness];  | 
|
299  | 
||
| 7031 | 300  | 
Goal "A <= UNIV";  | 
301  | 
by (rtac subsetI 1);  | 
|
302  | 
by (rtac UNIV_I 1);  | 
|
303  | 
qed "subset_UNIV";  | 
|
| 
4159
 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 
paulson 
parents: 
4135 
diff
changeset
 | 
304  | 
|
| 
 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 
paulson 
parents: 
4135 
diff
changeset
 | 
305  | 
(** Eta-contracting these two rules (to remove P) causes them to be ignored  | 
| 
 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 
paulson 
parents: 
4135 
diff
changeset
 | 
306  | 
because of their interaction with congruence rules. **)  | 
| 
 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 
paulson 
parents: 
4135 
diff
changeset
 | 
307  | 
|
| 5069 | 308  | 
Goalw [Ball_def] "Ball UNIV P = All P";  | 
| 
4159
 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 
paulson 
parents: 
4135 
diff
changeset
 | 
309  | 
by (Simp_tac 1);  | 
| 
 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 
paulson 
parents: 
4135 
diff
changeset
 | 
310  | 
qed "ball_UNIV";  | 
| 
 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 
paulson 
parents: 
4135 
diff
changeset
 | 
311  | 
|
| 5069 | 312  | 
Goalw [Bex_def] "Bex UNIV P = Ex P";  | 
| 
4159
 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 
paulson 
parents: 
4135 
diff
changeset
 | 
313  | 
by (Simp_tac 1);  | 
| 
 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 
paulson 
parents: 
4135 
diff
changeset
 | 
314  | 
qed "bex_UNIV";  | 
| 
 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 
paulson 
parents: 
4135 
diff
changeset
 | 
315  | 
Addsimps [ball_UNIV, bex_UNIV];  | 
| 
 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 
paulson 
parents: 
4135 
diff
changeset
 | 
316  | 
|
| 
 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 
paulson 
parents: 
4135 
diff
changeset
 | 
317  | 
|
| 2858 | 318  | 
section "The empty set -- {}";
 | 
319  | 
||
| 7007 | 320  | 
Goalw [empty_def] "(c : {}) = False";
 | 
321  | 
by (Blast_tac 1) ;  | 
|
322  | 
qed "empty_iff";  | 
|
| 2858 | 323  | 
|
324  | 
Addsimps [empty_iff];  | 
|
325  | 
||
| 7007 | 326  | 
Goal "a:{} ==> P";
 | 
327  | 
by (Full_simp_tac 1);  | 
|
328  | 
qed "emptyE";  | 
|
| 2858 | 329  | 
|
330  | 
AddSEs [emptyE];  | 
|
331  | 
||
| 7007 | 332  | 
Goal "{} <= A";
 | 
333  | 
by (Blast_tac 1) ;  | 
|
334  | 
qed "empty_subsetI";  | 
|
| 2858 | 335  | 
|
| 5256 | 336  | 
(*One effect is to delete the ASSUMPTION {} <= A*)
 | 
337  | 
AddIffs [empty_subsetI];  | 
|
338  | 
||
| 7031 | 339  | 
val [prem]= Goal "[| !!y. y:A ==> False |] ==> A={}";
 | 
| 7007 | 340  | 
by (blast_tac (claset() addIs [prem RS FalseE]) 1) ;  | 
341  | 
qed "equals0I";  | 
|
| 2858 | 342  | 
|
| 5256 | 343  | 
(*Use for reasoning about disjointness: A Int B = {} *)
 | 
| 7007 | 344  | 
Goal "A={} ==> a ~: A";
 | 
345  | 
by (Blast_tac 1) ;  | 
|
346  | 
qed "equals0D";  | 
|
| 2858 | 347  | 
|
| 5069 | 348  | 
Goalw [Ball_def] "Ball {} P = True";
 | 
| 
4159
 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 
paulson 
parents: 
4135 
diff
changeset
 | 
349  | 
by (Simp_tac 1);  | 
| 
 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 
paulson 
parents: 
4135 
diff
changeset
 | 
350  | 
qed "ball_empty";  | 
| 
 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 
paulson 
parents: 
4135 
diff
changeset
 | 
351  | 
|
| 5069 | 352  | 
Goalw [Bex_def] "Bex {} P = False";
 | 
| 
4159
 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 
paulson 
parents: 
4135 
diff
changeset
 | 
353  | 
by (Simp_tac 1);  | 
| 
 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 
paulson 
parents: 
4135 
diff
changeset
 | 
354  | 
qed "bex_empty";  | 
| 
 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 
paulson 
parents: 
4135 
diff
changeset
 | 
355  | 
Addsimps [ball_empty, bex_empty];  | 
| 
 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 
paulson 
parents: 
4135 
diff
changeset
 | 
356  | 
|
| 5069 | 357  | 
Goal "UNIV ~= {}";
 | 
| 
4159
 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 
paulson 
parents: 
4135 
diff
changeset
 | 
358  | 
by (blast_tac (claset() addEs [equalityE]) 1);  | 
| 
 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 
paulson 
parents: 
4135 
diff
changeset
 | 
359  | 
qed "UNIV_not_empty";  | 
| 
 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 
paulson 
parents: 
4135 
diff
changeset
 | 
360  | 
AddIffs [UNIV_not_empty];  | 
| 
 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 
paulson 
parents: 
4135 
diff
changeset
 | 
361  | 
|
| 
 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 
paulson 
parents: 
4135 
diff
changeset
 | 
362  | 
|
| 2858 | 363  | 
|
364  | 
section "The Powerset operator -- Pow";  | 
|
365  | 
||
| 7007 | 366  | 
Goalw [Pow_def] "(A : Pow(B)) = (A <= B)";  | 
367  | 
by (Asm_simp_tac 1);  | 
|
368  | 
qed "Pow_iff";  | 
|
| 2858 | 369  | 
|
370  | 
AddIffs [Pow_iff];  | 
|
371  | 
||
| 7031 | 372  | 
Goalw [Pow_def] "A <= B ==> A : Pow(B)";  | 
| 7007 | 373  | 
by (etac CollectI 1);  | 
374  | 
qed "PowI";  | 
|
| 2858 | 375  | 
|
| 7031 | 376  | 
Goalw [Pow_def] "A : Pow(B) ==> A<=B";  | 
| 7007 | 377  | 
by (etac CollectD 1);  | 
378  | 
qed "PowD";  | 
|
379  | 
||
| 2858 | 380  | 
|
| 9108 | 381  | 
bind_thm ("Pow_bottom", empty_subsetI RS PowI);        (* {}: Pow(B) *)
 | 
382  | 
bind_thm ("Pow_top", subset_refl RS PowI);             (* A : Pow(A) *)
 | 
|
| 2858 | 383  | 
|
384  | 
||
| 5931 | 385  | 
section "Set complement";  | 
| 923 | 386  | 
|
| 7031 | 387  | 
Goalw [Compl_def] "(c : -A) = (c~:A)";  | 
388  | 
by (Blast_tac 1);  | 
|
389  | 
qed "Compl_iff";  | 
|
| 
2499
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
390  | 
|
| 
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
391  | 
Addsimps [Compl_iff];  | 
| 
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
392  | 
|
| 5490 | 393  | 
val prems = Goalw [Compl_def] "[| c:A ==> False |] ==> c : -A";  | 
| 923 | 394  | 
by (REPEAT (ares_tac (prems @ [CollectI,notI]) 1));  | 
395  | 
qed "ComplI";  | 
|
396  | 
||
397  | 
(*This form, with negated conclusion, works well with the Classical prover.  | 
|
398  | 
Negated assumptions behave like formulae on the right side of the notional  | 
|
399  | 
turnstile...*)  | 
|
| 5490 | 400  | 
Goalw [Compl_def] "c : -A ==> c~:A";  | 
| 5316 | 401  | 
by (etac CollectD 1);  | 
| 923 | 402  | 
qed "ComplD";  | 
403  | 
||
| 9108 | 404  | 
bind_thm ("ComplE", make_elim ComplD);
 | 
| 923 | 405  | 
|
| 
2499
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
406  | 
AddSIs [ComplI];  | 
| 
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
407  | 
AddSEs [ComplE];  | 
| 1640 | 408  | 
|
| 923 | 409  | 
|
| 1548 | 410  | 
section "Binary union -- Un";  | 
| 923 | 411  | 
|
| 7031 | 412  | 
Goalw [Un_def] "(c : A Un B) = (c:A | c:B)";  | 
413  | 
by (Blast_tac 1);  | 
|
414  | 
qed "Un_iff";  | 
|
| 
2499
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
415  | 
Addsimps [Un_iff];  | 
| 
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
416  | 
|
| 
5143
 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 
paulson 
parents: 
5069 
diff
changeset
 | 
417  | 
Goal "c:A ==> c : A Un B";  | 
| 
2499
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
418  | 
by (Asm_simp_tac 1);  | 
| 923 | 419  | 
qed "UnI1";  | 
420  | 
||
| 
5143
 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 
paulson 
parents: 
5069 
diff
changeset
 | 
421  | 
Goal "c:B ==> c : A Un B";  | 
| 
2499
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
422  | 
by (Asm_simp_tac 1);  | 
| 923 | 423  | 
qed "UnI2";  | 
424  | 
||
| 11589 | 425  | 
AddXEs [UnI1, UnI2];  | 
| 9378 | 426  | 
|
427  | 
||
| 923 | 428  | 
(*Classical introduction rule: no commitment to A vs B*)  | 
| 7007 | 429  | 
|
| 7031 | 430  | 
val prems = Goal "(c~:B ==> c:A) ==> c : A Un B";  | 
| 7007 | 431  | 
by (Simp_tac 1);  | 
432  | 
by (REPEAT (ares_tac (prems@[disjCI]) 1)) ;  | 
|
433  | 
qed "UnCI";  | 
|
| 923 | 434  | 
|
| 5316 | 435  | 
val major::prems = Goalw [Un_def]  | 
| 923 | 436  | 
"[| c : A Un B; c:A ==> P; c:B ==> P |] ==> P";  | 
437  | 
by (rtac (major RS CollectD RS disjE) 1);  | 
|
438  | 
by (REPEAT (eresolve_tac prems 1));  | 
|
439  | 
qed "UnE";  | 
|
440  | 
||
| 
2499
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
441  | 
AddSIs [UnCI];  | 
| 
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
442  | 
AddSEs [UnE];  | 
| 1640 | 443  | 
|
| 923 | 444  | 
|
| 1548 | 445  | 
section "Binary intersection -- Int";  | 
| 923 | 446  | 
|
| 7031 | 447  | 
Goalw [Int_def] "(c : A Int B) = (c:A & c:B)";  | 
448  | 
by (Blast_tac 1);  | 
|
449  | 
qed "Int_iff";  | 
|
| 
2499
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
450  | 
Addsimps [Int_iff];  | 
| 
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
451  | 
|
| 
5143
 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 
paulson 
parents: 
5069 
diff
changeset
 | 
452  | 
Goal "[| c:A; c:B |] ==> c : A Int B";  | 
| 
2499
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
453  | 
by (Asm_simp_tac 1);  | 
| 923 | 454  | 
qed "IntI";  | 
455  | 
||
| 
5143
 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 
paulson 
parents: 
5069 
diff
changeset
 | 
456  | 
Goal "c : A Int B ==> c:A";  | 
| 
2499
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
457  | 
by (Asm_full_simp_tac 1);  | 
| 923 | 458  | 
qed "IntD1";  | 
459  | 
||
| 
5143
 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 
paulson 
parents: 
5069 
diff
changeset
 | 
460  | 
Goal "c : A Int B ==> c:B";  | 
| 
2499
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
461  | 
by (Asm_full_simp_tac 1);  | 
| 923 | 462  | 
qed "IntD2";  | 
463  | 
||
| 5316 | 464  | 
val [major,minor] = Goal  | 
| 923 | 465  | 
"[| c : A Int B; [| c:A; c:B |] ==> P |] ==> P";  | 
466  | 
by (rtac minor 1);  | 
|
467  | 
by (rtac (major RS IntD1) 1);  | 
|
468  | 
by (rtac (major RS IntD2) 1);  | 
|
469  | 
qed "IntE";  | 
|
470  | 
||
| 
2499
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
471  | 
AddSIs [IntI];  | 
| 
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
472  | 
AddSEs [IntE];  | 
| 923 | 473  | 
|
| 1548 | 474  | 
section "Set difference";  | 
| 923 | 475  | 
|
| 7031 | 476  | 
Goalw [set_diff_def] "(c : A-B) = (c:A & c~:B)";  | 
477  | 
by (Blast_tac 1);  | 
|
478  | 
qed "Diff_iff";  | 
|
| 
2499
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
479  | 
Addsimps [Diff_iff];  | 
| 
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
480  | 
|
| 7007 | 481  | 
Goal "[| c : A; c ~: B |] ==> c : A - B";  | 
482  | 
by (Asm_simp_tac 1) ;  | 
|
483  | 
qed "DiffI";  | 
|
| 923 | 484  | 
|
| 7007 | 485  | 
Goal "c : A - B ==> c : A";  | 
486  | 
by (Asm_full_simp_tac 1) ;  | 
|
487  | 
qed "DiffD1";  | 
|
| 923 | 488  | 
|
| 7007 | 489  | 
Goal "[| c : A - B; c : B |] ==> P";  | 
490  | 
by (Asm_full_simp_tac 1) ;  | 
|
491  | 
qed "DiffD2";  | 
|
| 
2499
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
492  | 
|
| 7031 | 493  | 
val prems = Goal "[| c : A - B; [| c:A; c~:B |] ==> P |] ==> P";  | 
| 7007 | 494  | 
by (resolve_tac prems 1);  | 
495  | 
by (REPEAT (ares_tac (prems RL [DiffD1, DiffD2 RS notI]) 1)) ;  | 
|
496  | 
qed "DiffE";  | 
|
| 923 | 497  | 
|
| 
2499
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
498  | 
AddSIs [DiffI];  | 
| 
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
499  | 
AddSEs [DiffE];  | 
| 923 | 500  | 
|
501  | 
||
| 1548 | 502  | 
section "Augmenting a set -- insert";  | 
| 923 | 503  | 
|
| 11655 | 504  | 
Goalw [insert_def] "(a : insert b A) = (a=b | a:A)";  | 
| 7031 | 505  | 
by (Blast_tac 1);  | 
506  | 
qed "insert_iff";  | 
|
| 
2499
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
507  | 
Addsimps [insert_iff];  | 
| 923 | 508  | 
|
| 7031 | 509  | 
Goal "a : insert a B";  | 
| 7007 | 510  | 
by (Simp_tac 1);  | 
511  | 
qed "insertI1";  | 
|
| 923 | 512  | 
|
| 7007 | 513  | 
Goal "!!a. a : B ==> a : insert b B";  | 
514  | 
by (Asm_simp_tac 1);  | 
|
515  | 
qed "insertI2";  | 
|
516  | 
||
517  | 
val major::prems = Goalw [insert_def]  | 
|
518  | 
"[| a : insert b A; a=b ==> P; a:A ==> P |] ==> P";  | 
|
519  | 
by (rtac (major RS UnE) 1);  | 
|
520  | 
by (REPEAT (eresolve_tac (prems @ [CollectE]) 1));  | 
|
521  | 
qed "insertE";  | 
|
| 923 | 522  | 
|
523  | 
(*Classical introduction rule*)  | 
|
| 7031 | 524  | 
val prems = Goal "(a~:B ==> a=b) ==> a: insert b B";  | 
| 7007 | 525  | 
by (Simp_tac 1);  | 
526  | 
by (REPEAT (ares_tac (prems@[disjCI]) 1)) ;  | 
|
527  | 
qed "insertCI";  | 
|
| 
2499
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
528  | 
|
| 
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
529  | 
AddSIs [insertCI];  | 
| 
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
530  | 
AddSEs [insertE];  | 
| 923 | 531  | 
|
| 
9088
 
453996655ac2
replaced the useless [p]subset_insertD by [p]subset_insert_iff
 
paulson 
parents: 
9075 
diff
changeset
 | 
532  | 
Goal "(A <= insert x B) = (if x:A then A-{x} <= B else A<=B)";
 | 
| 
 
453996655ac2
replaced the useless [p]subset_insertD by [p]subset_insert_iff
 
paulson 
parents: 
9075 
diff
changeset
 | 
533  | 
by Auto_tac;  | 
| 
 
453996655ac2
replaced the useless [p]subset_insertD by [p]subset_insert_iff
 
paulson 
parents: 
9075 
diff
changeset
 | 
534  | 
qed "subset_insert_iff";  | 
| 
7496
 
93ae11d887ff
added theorems subset_insertD, singleton_insert_inj_eq, subset_singletonD
 
oheimb 
parents: 
7441 
diff
changeset
 | 
535  | 
|
| 1548 | 536  | 
section "Singletons, using insert";  | 
| 923 | 537  | 
|
| 7007 | 538  | 
Goal "a : {a}";
 | 
539  | 
by (rtac insertI1 1) ;  | 
|
540  | 
qed "singletonI";  | 
|
| 923 | 541  | 
|
| 
5143
 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 
paulson 
parents: 
5069 
diff
changeset
 | 
542  | 
Goal "b : {a} ==> b=a";
 | 
| 2891 | 543  | 
by (Blast_tac 1);  | 
| 923 | 544  | 
qed "singletonD";  | 
545  | 
||
| 
1776
 
d7e77cb8ce5c
moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
 
oheimb 
parents: 
1762 
diff
changeset
 | 
546  | 
bind_thm ("singletonE", make_elim singletonD);
 | 
| 
 
d7e77cb8ce5c
moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
 
oheimb 
parents: 
1762 
diff
changeset
 | 
547  | 
|
| 7007 | 548  | 
Goal "(b : {a}) = (b=a)";
 | 
549  | 
by (Blast_tac 1);  | 
|
550  | 
qed "singleton_iff";  | 
|
| 923 | 551  | 
|
| 
5143
 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 
paulson 
parents: 
5069 
diff
changeset
 | 
552  | 
Goal "{a}={b} ==> a=b";
 | 
| 4089 | 553  | 
by (blast_tac (claset() addEs [equalityE]) 1);  | 
| 923 | 554  | 
qed "singleton_inject";  | 
555  | 
||
| 2858 | 556  | 
(*Redundant? But unlike insertCI, it proves the subgoal immediately!*)  | 
557  | 
AddSIs [singletonI];  | 
|
| 
2499
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
558  | 
AddSDs [singleton_inject];  | 
| 3718 | 559  | 
AddSEs [singletonE];  | 
| 
2499
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
560  | 
|
| 11655 | 561  | 
Goal "({b} = insert a A) = (a = b & A <= {b})";
 | 
| 
8326
 
0e329578b0ef
tidied the proofs of singleton_insert_inj_eq, singleton_insert_inj_eq' and
 
paulson 
parents: 
8053 
diff
changeset
 | 
562  | 
by (blast_tac (claset() addSEs [equalityE]) 1);  | 
| 
7496
 
93ae11d887ff
added theorems subset_insertD, singleton_insert_inj_eq, subset_singletonD
 
oheimb 
parents: 
7441 
diff
changeset
 | 
563  | 
qed "singleton_insert_inj_eq";  | 
| 
 
93ae11d887ff
added theorems subset_insertD, singleton_insert_inj_eq, subset_singletonD
 
oheimb 
parents: 
7441 
diff
changeset
 | 
564  | 
|
| 
8326
 
0e329578b0ef
tidied the proofs of singleton_insert_inj_eq, singleton_insert_inj_eq' and
 
paulson 
parents: 
8053 
diff
changeset
 | 
565  | 
Goal "(insert a A = {b}) = (a = b & A <= {b})";
 | 
| 
 
0e329578b0ef
tidied the proofs of singleton_insert_inj_eq, singleton_insert_inj_eq' and
 
paulson 
parents: 
8053 
diff
changeset
 | 
566  | 
by (blast_tac (claset() addSEs [equalityE]) 1);  | 
| 7969 | 567  | 
qed "singleton_insert_inj_eq'";  | 
568  | 
||
| 
8326
 
0e329578b0ef
tidied the proofs of singleton_insert_inj_eq, singleton_insert_inj_eq' and
 
paulson 
parents: 
8053 
diff
changeset
 | 
569  | 
AddIffs [singleton_insert_inj_eq, singleton_insert_inj_eq'];  | 
| 
 
0e329578b0ef
tidied the proofs of singleton_insert_inj_eq, singleton_insert_inj_eq' and
 
paulson 
parents: 
8053 
diff
changeset
 | 
570  | 
|
| 
7496
 
93ae11d887ff
added theorems subset_insertD, singleton_insert_inj_eq, subset_singletonD
 
oheimb 
parents: 
7441 
diff
changeset
 | 
571  | 
Goal "A <= {x} ==> A={} | A = {x}";
 | 
| 
 
93ae11d887ff
added theorems subset_insertD, singleton_insert_inj_eq, subset_singletonD
 
oheimb 
parents: 
7441 
diff
changeset
 | 
572  | 
by (Fast_tac 1);  | 
| 
 
93ae11d887ff
added theorems subset_insertD, singleton_insert_inj_eq, subset_singletonD
 
oheimb 
parents: 
7441 
diff
changeset
 | 
573  | 
qed "subset_singletonD";  | 
| 
 
93ae11d887ff
added theorems subset_insertD, singleton_insert_inj_eq, subset_singletonD
 
oheimb 
parents: 
7441 
diff
changeset
 | 
574  | 
|
| 5069 | 575  | 
Goal "{x. x=a} = {a}";
 | 
| 4423 | 576  | 
by (Blast_tac 1);  | 
| 3582 | 577  | 
qed "singleton_conv";  | 
578  | 
Addsimps [singleton_conv];  | 
|
| 1531 | 579  | 
|
| 5600 | 580  | 
Goal "{x. a=x} = {a}";
 | 
| 6301 | 581  | 
by (Blast_tac 1);  | 
| 5600 | 582  | 
qed "singleton_conv2";  | 
583  | 
Addsimps [singleton_conv2];  | 
|
584  | 
||
| 11007 | 585  | 
Goal "[| A - {x} <= B; x : A |] ==> A <= insert x B"; 
 | 
586  | 
by(Blast_tac 1);  | 
|
587  | 
qed "diff_single_insert";  | 
|
588  | 
||
| 1531 | 589  | 
|
| 10832 | 590  | 
section "Unions of families -- UNION x:A. B(x) is Union(B`A)";  | 
| 923 | 591  | 
|
| 5069 | 592  | 
Goalw [UNION_def] "(b: (UN x:A. B(x))) = (EX x:A. b: B(x))";  | 
| 2891 | 593  | 
by (Blast_tac 1);  | 
| 
2499
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
594  | 
qed "UN_iff";  | 
| 
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
595  | 
|
| 
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
596  | 
Addsimps [UN_iff];  | 
| 
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
597  | 
|
| 923 | 598  | 
(*The order of the premises presupposes that A is rigid; b may be flexible*)  | 
| 
5143
 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 
paulson 
parents: 
5069 
diff
changeset
 | 
599  | 
Goal "[| a:A; b: B(a) |] ==> b: (UN x:A. B(x))";  | 
| 
4477
 
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
 
paulson 
parents: 
4469 
diff
changeset
 | 
600  | 
by Auto_tac;  | 
| 923 | 601  | 
qed "UN_I";  | 
602  | 
||
| 5316 | 603  | 
val major::prems = Goalw [UNION_def]  | 
| 923 | 604  | 
"[| b : (UN x:A. B(x)); !!x.[| x:A; b: B(x) |] ==> R |] ==> R";  | 
605  | 
by (rtac (major RS CollectD RS bexE) 1);  | 
|
606  | 
by (REPEAT (ares_tac prems 1));  | 
|
607  | 
qed "UN_E";  | 
|
608  | 
||
| 
2499
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
609  | 
AddIs [UN_I];  | 
| 
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
610  | 
AddSEs [UN_E];  | 
| 
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
611  | 
|
| 6291 | 612  | 
val prems = Goalw [UNION_def]  | 
| 923 | 613  | 
"[| A=B; !!x. x:B ==> C(x) = D(x) |] ==> \  | 
614  | 
\ (UN x:A. C(x)) = (UN x:B. D(x))";  | 
|
| 6291 | 615  | 
by (asm_simp_tac (simpset() addsimps prems) 1);  | 
| 923 | 616  | 
qed "UN_cong";  | 
| 
9687
 
772ac061bd76
moved congruence rules UN_cong, INT_cong from UNTIY/Union to Set.ML
 
paulson 
parents: 
9422 
diff
changeset
 | 
617  | 
Addcongs [UN_cong];  | 
| 923 | 618  | 
|
619  | 
||
| 10832 | 620  | 
section "Intersections of families -- INTER x:A. B(x) is Inter(B`A)";  | 
| 923 | 621  | 
|
| 5069 | 622  | 
Goalw [INTER_def] "(b: (INT x:A. B(x))) = (ALL x:A. b: B(x))";  | 
| 
4477
 
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
 
paulson 
parents: 
4469 
diff
changeset
 | 
623  | 
by Auto_tac;  | 
| 
2499
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
624  | 
qed "INT_iff";  | 
| 
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
625  | 
|
| 
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
626  | 
Addsimps [INT_iff];  | 
| 
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
627  | 
|
| 5316 | 628  | 
val prems = Goalw [INTER_def]  | 
| 923 | 629  | 
"(!!x. x:A ==> b: B(x)) ==> b : (INT x:A. B(x))";  | 
630  | 
by (REPEAT (ares_tac ([CollectI,ballI] @ prems) 1));  | 
|
631  | 
qed "INT_I";  | 
|
632  | 
||
| 
5143
 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 
paulson 
parents: 
5069 
diff
changeset
 | 
633  | 
Goal "[| b : (INT x:A. B(x)); a:A |] ==> b: B(a)";  | 
| 
4477
 
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
 
paulson 
parents: 
4469 
diff
changeset
 | 
634  | 
by Auto_tac;  | 
| 923 | 635  | 
qed "INT_D";  | 
636  | 
||
637  | 
(*"Classical" elimination -- by the Excluded Middle on a:A *)  | 
|
| 5316 | 638  | 
val major::prems = Goalw [INTER_def]  | 
| 923 | 639  | 
"[| b : (INT x:A. B(x)); b: B(a) ==> R; a~:A ==> R |] ==> R";  | 
640  | 
by (rtac (major RS CollectD RS ballE) 1);  | 
|
641  | 
by (REPEAT (eresolve_tac prems 1));  | 
|
642  | 
qed "INT_E";  | 
|
643  | 
||
| 
2499
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
644  | 
AddSIs [INT_I];  | 
| 
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
645  | 
AddEs [INT_D, INT_E];  | 
| 
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
646  | 
|
| 6291 | 647  | 
val prems = Goalw [INTER_def]  | 
| 923 | 648  | 
"[| A=B; !!x. x:B ==> C(x) = D(x) |] ==> \  | 
649  | 
\ (INT x:A. C(x)) = (INT x:B. D(x))";  | 
|
| 6291 | 650  | 
by (asm_simp_tac (simpset() addsimps prems) 1);  | 
| 923 | 651  | 
qed "INT_cong";  | 
| 
9687
 
772ac061bd76
moved congruence rules UN_cong, INT_cong from UNTIY/Union to Set.ML
 
paulson 
parents: 
9422 
diff
changeset
 | 
652  | 
Addcongs [INT_cong];  | 
| 923 | 653  | 
|
654  | 
||
| 1548 | 655  | 
section "Union";  | 
| 923 | 656  | 
|
| 5069 | 657  | 
Goalw [Union_def] "(A : Union(C)) = (EX X:C. A:X)";  | 
| 2891 | 658  | 
by (Blast_tac 1);  | 
| 
2499
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
659  | 
qed "Union_iff";  | 
| 
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
660  | 
|
| 
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
661  | 
Addsimps [Union_iff];  | 
| 
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
662  | 
|
| 923 | 663  | 
(*The order of the premises presupposes that C is rigid; A may be flexible*)  | 
| 
5143
 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 
paulson 
parents: 
5069 
diff
changeset
 | 
664  | 
Goal "[| X:C; A:X |] ==> A : Union(C)";  | 
| 
4477
 
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
 
paulson 
parents: 
4469 
diff
changeset
 | 
665  | 
by Auto_tac;  | 
| 923 | 666  | 
qed "UnionI";  | 
667  | 
||
| 5316 | 668  | 
val major::prems = Goalw [Union_def]  | 
| 923 | 669  | 
"[| A : Union(C); !!X.[| A:X; X:C |] ==> R |] ==> R";  | 
670  | 
by (rtac (major RS UN_E) 1);  | 
|
671  | 
by (REPEAT (ares_tac prems 1));  | 
|
672  | 
qed "UnionE";  | 
|
673  | 
||
| 
2499
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
674  | 
AddIs [UnionI];  | 
| 
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
675  | 
AddSEs [UnionE];  | 
| 
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
676  | 
|
| 
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
677  | 
|
| 1548 | 678  | 
section "Inter";  | 
| 923 | 679  | 
|
| 5069 | 680  | 
Goalw [Inter_def] "(A : Inter(C)) = (ALL X:C. A:X)";  | 
| 2891 | 681  | 
by (Blast_tac 1);  | 
| 
2499
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
682  | 
qed "Inter_iff";  | 
| 
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
683  | 
|
| 
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
684  | 
Addsimps [Inter_iff];  | 
| 
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
685  | 
|
| 5316 | 686  | 
val prems = Goalw [Inter_def]  | 
| 923 | 687  | 
"[| !!X. X:C ==> A:X |] ==> A : Inter(C)";  | 
688  | 
by (REPEAT (ares_tac ([INT_I] @ prems) 1));  | 
|
689  | 
qed "InterI";  | 
|
690  | 
||
691  | 
(*A "destruct" rule -- every X in C contains A as an element, but  | 
|
692  | 
A:X can hold when X:C does not! This rule is analogous to "spec". *)  | 
|
| 
5143
 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 
paulson 
parents: 
5069 
diff
changeset
 | 
693  | 
Goal "[| A : Inter(C); X:C |] ==> A:X";  | 
| 
4477
 
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
 
paulson 
parents: 
4469 
diff
changeset
 | 
694  | 
by Auto_tac;  | 
| 923 | 695  | 
qed "InterD";  | 
696  | 
||
697  | 
(*"Classical" elimination rule -- does not require proving X:C *)  | 
|
| 5316 | 698  | 
val major::prems = Goalw [Inter_def]  | 
| 2721 | 699  | 
"[| A : Inter(C); X~:C ==> R; A:X ==> R |] ==> R";  | 
| 923 | 700  | 
by (rtac (major RS INT_E) 1);  | 
701  | 
by (REPEAT (eresolve_tac prems 1));  | 
|
702  | 
qed "InterE";  | 
|
703  | 
||
| 
2499
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
704  | 
AddSIs [InterI];  | 
| 
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
705  | 
AddEs [InterD, InterE];  | 
| 
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
706  | 
|
| 
 
0bc87b063447
Tidying of proofs.  New theorems are enterred immediately into the
 
paulson 
parents: 
2031 
diff
changeset
 | 
707  | 
|
| 2912 | 708  | 
(*** Image of a set under a function ***)  | 
709  | 
||
710  | 
(*Frequently b does not have the syntactic form of f(x).*)  | 
|
| 10832 | 711  | 
Goalw [image_def] "[| b=f(x); x:A |] ==> b : f`A";  | 
| 5316 | 712  | 
by (Blast_tac 1);  | 
| 2912 | 713  | 
qed "image_eqI";  | 
| 3909 | 714  | 
Addsimps [image_eqI];  | 
| 2912 | 715  | 
|
716  | 
bind_thm ("imageI", refl RS image_eqI);
 | 
|
717  | 
||
| 8025 | 718  | 
(*This version's more effective when we already have the required x*)  | 
| 10832 | 719  | 
Goalw [image_def] "[| x:A; b=f(x) |] ==> b : f`A";  | 
| 8025 | 720  | 
by (Blast_tac 1);  | 
721  | 
qed "rev_image_eqI";  | 
|
722  | 
||
| 2912 | 723  | 
(*The eta-expansion gives variable-name preservation.*)  | 
| 5316 | 724  | 
val major::prems = Goalw [image_def]  | 
| 10832 | 725  | 
"[| b : (%x. f(x))`A; !!x.[| b=f(x); x:A |] ==> P |] ==> P";  | 
| 2912 | 726  | 
by (rtac (major RS CollectD RS bexE) 1);  | 
727  | 
by (REPEAT (ares_tac prems 1));  | 
|
728  | 
qed "imageE";  | 
|
729  | 
||
730  | 
AddIs [image_eqI];  | 
|
731  | 
AddSEs [imageE];  | 
|
732  | 
||
| 10832 | 733  | 
Goal "f`(A Un B) = f`A Un f`B";  | 
| 2935 | 734  | 
by (Blast_tac 1);  | 
| 2912 | 735  | 
qed "image_Un";  | 
736  | 
||
| 10832 | 737  | 
Goal "(z : f`A) = (EX x:A. z = f x)";  | 
| 3960 | 738  | 
by (Blast_tac 1);  | 
739  | 
qed "image_iff";  | 
|
740  | 
||
| 4523 | 741  | 
(*This rewrite rule would confuse users if made default.*)  | 
| 10832 | 742  | 
Goal "(f`A <= B) = (ALL x:A. f(x): B)";  | 
| 4523 | 743  | 
by (Blast_tac 1);  | 
744  | 
qed "image_subset_iff";  | 
|
745  | 
||
| 11655 | 746  | 
Goal "(B <= f ` A) = (? AA. AA <= A & B = f ` AA)";  | 
| 11007 | 747  | 
by Safe_tac;  | 
748  | 
by (Fast_tac 2);  | 
|
749  | 
by (res_inst_tac [("x","{a. a : A & f a : B}")] exI 1);
 | 
|
750  | 
by (Fast_tac 1);  | 
|
751  | 
qed "subset_image_iff";  | 
|
752  | 
||
| 4523 | 753  | 
(*Replaces the three steps subsetI, imageE, hyp_subst_tac, but breaks too  | 
754  | 
many existing proofs.*)  | 
|
| 10832 | 755  | 
val prems = Goal "(!!x. x:A ==> f(x) : B) ==> f`A <= B";  | 
| 4510 | 756  | 
by (blast_tac (claset() addIs prems) 1);  | 
757  | 
qed "image_subsetI";  | 
|
758  | 
||
| 2912 | 759  | 
(*** Range of a function -- just a translation for image! ***)  | 
760  | 
||
| 
5143
 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 
paulson 
parents: 
5069 
diff
changeset
 | 
761  | 
Goal "b=f(x) ==> b : range(f)";  | 
| 2912 | 762  | 
by (EVERY1 [etac image_eqI, rtac UNIV_I]);  | 
763  | 
bind_thm ("range_eqI", UNIV_I RSN (2,image_eqI));
 | 
|
764  | 
||
765  | 
bind_thm ("rangeI", UNIV_I RS imageI);
 | 
|
766  | 
||
| 5316 | 767  | 
val [major,minor] = Goal  | 
| 3842 | 768  | 
"[| b : range(%x. f(x)); !!x. b=f(x) ==> P |] ==> P";  | 
| 2912 | 769  | 
by (rtac (major RS imageE) 1);  | 
770  | 
by (etac minor 1);  | 
|
771  | 
qed "rangeE";  | 
|
| 10482 | 772  | 
AddXEs [rangeE];  | 
| 2912 | 773  | 
|
| 
1776
 
d7e77cb8ce5c
moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
 
oheimb 
parents: 
1762 
diff
changeset
 | 
774  | 
|
| 
 
d7e77cb8ce5c
moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
 
oheimb 
parents: 
1762 
diff
changeset
 | 
775  | 
(*** Set reasoning tools ***)  | 
| 
 
d7e77cb8ce5c
moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
 
oheimb 
parents: 
1762 
diff
changeset
 | 
776  | 
|
| 
 
d7e77cb8ce5c
moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
 
oheimb 
parents: 
1762 
diff
changeset
 | 
777  | 
|
| 3912 | 778  | 
(** Rewrite rules for boolean case-splitting: faster than  | 
| 4830 | 779  | 
addsplits[split_if]  | 
| 3912 | 780  | 
**)  | 
781  | 
||
| 4830 | 782  | 
bind_thm ("split_if_eq1", read_instantiate [("P", "%x. x = ?b")] split_if);
 | 
783  | 
bind_thm ("split_if_eq2", read_instantiate [("P", "%x. ?a = x")] split_if);
 | 
|
| 3912 | 784  | 
|
| 5237 | 785  | 
(*Split ifs on either side of the membership relation.  | 
786  | 
Not for Addsimps -- can cause goals to blow up!*)  | 
|
| 9969 | 787  | 
bind_thm ("split_if_mem1", inst "P" "%x. x : ?b" split_if);
 | 
788  | 
bind_thm ("split_if_mem2", inst "P" "%x. ?a : x" split_if);
 | 
|
| 3912 | 789  | 
|
| 9108 | 790  | 
bind_thms ("split_ifs", [if_bool_eq_conj, split_if_eq1, split_if_eq2,
 | 
| 9969 | 791  | 
split_if_mem1, split_if_mem2]);  | 
| 3912 | 792  | 
|
793  | 
||
| 4089 | 794  | 
(*Each of these has ALREADY been added to simpset() above.*)  | 
| 9108 | 795  | 
bind_thms ("mem_simps", [insert_iff, empty_iff, Un_iff, Int_iff, Compl_iff, Diff_iff, 
 | 
796  | 
mem_Collect_eq, UN_iff, Union_iff, INT_iff, Inter_iff]);  | 
|
| 
1776
 
d7e77cb8ce5c
moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
 
oheimb 
parents: 
1762 
diff
changeset
 | 
797  | 
|
| 9041 | 798  | 
(*Would like to add these, but the existing code only searches for the  | 
799  | 
outer-level constant, which in this case is just "op :"; we instead need  | 
|
800  | 
to use term-nets to associate patterns with rules. Also, if a rule fails to  | 
|
801  | 
apply, then the formula should be kept.  | 
|
802  | 
  [("uminus", Compl_iff RS iffD1), ("op -", [Diff_iff RS iffD1]), 
 | 
|
803  | 
   ("op Int", [IntD1,IntD2]),
 | 
|
804  | 
   ("Collect", [CollectD]), ("Inter", [InterD]), ("INTER", [INT_D])]
 | 
|
805  | 
*)  | 
|
806  | 
val mksimps_pairs =  | 
|
807  | 
  [("Ball",[bspec])] @ mksimps_pairs;
 | 
|
| 
1776
 
d7e77cb8ce5c
moved mem_simps and the corresponding update of !simpset from Fun.ML to Set.ML,
 
oheimb 
parents: 
1762 
diff
changeset
 | 
808  | 
|
| 6291 | 809  | 
simpset_ref() := simpset() setmksimps (mksimps mksimps_pairs);  | 
| 
3222
 
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
 
nipkow 
parents: 
2935 
diff
changeset
 | 
810  | 
|
| 5256 | 811  | 
Addsimps[subset_UNIV, subset_refl];  | 
| 
3222
 
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
 
nipkow 
parents: 
2935 
diff
changeset
 | 
812  | 
|
| 
 
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
 
nipkow 
parents: 
2935 
diff
changeset
 | 
813  | 
|
| 8001 | 814  | 
(*** The 'proper subset' relation (<) ***)  | 
| 
3222
 
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
 
nipkow 
parents: 
2935 
diff
changeset
 | 
815  | 
|
| 5069 | 816  | 
Goalw [psubset_def] "!!A::'a set. [| A <= B; A ~= B |] ==> A<B";  | 
| 
3222
 
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
 
nipkow 
parents: 
2935 
diff
changeset
 | 
817  | 
by (Blast_tac 1);  | 
| 
 
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
 
nipkow 
parents: 
2935 
diff
changeset
 | 
818  | 
qed "psubsetI";  | 
| 8913 | 819  | 
AddSIs [psubsetI];  | 
| 
3222
 
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
 
nipkow 
parents: 
2935 
diff
changeset
 | 
820  | 
|
| 
9088
 
453996655ac2
replaced the useless [p]subset_insertD by [p]subset_insert_iff
 
paulson 
parents: 
9075 
diff
changeset
 | 
821  | 
Goalw [psubset_def]  | 
| 
 
453996655ac2
replaced the useless [p]subset_insertD by [p]subset_insert_iff
 
paulson 
parents: 
9075 
diff
changeset
 | 
822  | 
  "(A < insert x B) = (if x:B then A<B else if x:A then A-{x} < B else A<=B)";
 | 
| 
 
453996655ac2
replaced the useless [p]subset_insertD by [p]subset_insert_iff
 
paulson 
parents: 
9075 
diff
changeset
 | 
823  | 
by (asm_simp_tac (simpset() addsimps [subset_insert_iff]) 1);  | 
| 
 
453996655ac2
replaced the useless [p]subset_insertD by [p]subset_insert_iff
 
paulson 
parents: 
9075 
diff
changeset
 | 
824  | 
by (Blast_tac 1);  | 
| 
 
453996655ac2
replaced the useless [p]subset_insertD by [p]subset_insert_iff
 
paulson 
parents: 
9075 
diff
changeset
 | 
825  | 
qed "psubset_insert_iff";  | 
| 4059 | 826  | 
|
827  | 
bind_thm ("psubset_eq", psubset_def RS meta_eq_to_obj_eq);
 | 
|
| 6443 | 828  | 
|
829  | 
bind_thm ("psubset_imp_subset", psubset_eq RS iffD1 RS conjunct1);
 | 
|
830  | 
||
831  | 
Goal"[| (A::'a set) < B; B <= C |] ==> A < C";  | 
|
832  | 
by (auto_tac (claset(), simpset() addsimps [psubset_eq]));  | 
|
833  | 
qed "psubset_subset_trans";  | 
|
834  | 
||
835  | 
Goal"[| (A::'a set) <= B; B < C|] ==> A < C";  | 
|
836  | 
by (auto_tac (claset(), simpset() addsimps [psubset_eq]));  | 
|
837  | 
qed "subset_psubset_trans";  | 
|
| 
7717
 
e7ecfa617443
Added attribute rulify_prems (useful for modifying premises of introduction
 
berghofe 
parents: 
7658 
diff
changeset
 | 
838  | 
|
| 8001 | 839  | 
Goalw [psubset_def] "A < B ==> EX b. b : (B - A)";  | 
840  | 
by (Blast_tac 1);  | 
|
841  | 
qed "psubset_imp_ex_mem";  | 
|
842  | 
||
| 
7717
 
e7ecfa617443
Added attribute rulify_prems (useful for modifying premises of introduction
 
berghofe 
parents: 
7658 
diff
changeset
 | 
843  | 
|
| 9892 | 844  | 
(* rulify setup *)  | 
845  | 
||
846  | 
Goal "(!!x. x:A ==> P x) == Trueprop (ALL x:A. P x)";  | 
|
847  | 
by (simp_tac (simpset () addsimps (Ball_def :: thms "atomize")) 1);  | 
|
848  | 
qed "ball_eq";  | 
|
| 
7717
 
e7ecfa617443
Added attribute rulify_prems (useful for modifying premises of introduction
 
berghofe 
parents: 
7658 
diff
changeset
 | 
849  | 
|
| 
 
e7ecfa617443
Added attribute rulify_prems (useful for modifying premises of introduction
 
berghofe 
parents: 
7658 
diff
changeset
 | 
850  | 
local  | 
| 9892 | 851  | 
val ss = HOL_basic_ss addsimps  | 
852  | 
(Drule.norm_hhf_eq :: map Thm.symmetric (ball_eq :: thms "atomize"));  | 
|
| 
7717
 
e7ecfa617443
Added attribute rulify_prems (useful for modifying premises of introduction
 
berghofe 
parents: 
7658 
diff
changeset
 | 
853  | 
in  | 
| 
 
e7ecfa617443
Added attribute rulify_prems (useful for modifying premises of introduction
 
berghofe 
parents: 
7658 
diff
changeset
 | 
854  | 
|
| 9892 | 855  | 
structure Rulify = RulifyFun  | 
856  | 
(val make_meta = Simplifier.simplify ss  | 
|
857  | 
val full_make_meta = Simplifier.full_simplify ss);  | 
|
858  | 
||
859  | 
structure BasicRulify: BASIC_RULIFY = Rulify;  | 
|
860  | 
open BasicRulify;  | 
|
| 
7717
 
e7ecfa617443
Added attribute rulify_prems (useful for modifying premises of introduction
 
berghofe 
parents: 
7658 
diff
changeset
 | 
861  | 
|
| 
 
e7ecfa617443
Added attribute rulify_prems (useful for modifying premises of introduction
 
berghofe 
parents: 
7658 
diff
changeset
 | 
862  | 
end;  |