| 1465 |      1 | (*  Title:      HOL/subset
 | 
| 923 |      2 |     ID:         $Id$
 | 
| 1465 |      3 |     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
 | 
| 923 |      4 |     Copyright   1991  University of Cambridge
 | 
|  |      5 | 
 | 
|  |      6 | Derived rules involving subsets
 | 
|  |      7 | Union and Intersection as lattice operations
 | 
|  |      8 | *)
 | 
|  |      9 | 
 | 
|  |     10 | (*** insert ***)
 | 
|  |     11 | 
 | 
|  |     12 | qed_goal "subset_insertI" Set.thy "B <= insert a B"
 | 
|  |     13 |  (fn _=> [ (rtac subsetI 1), (etac insertI2 1) ]);
 | 
|  |     14 | 
 | 
| 5316 |     15 | Goal "x ~: A ==> (A <= insert x B) = (A <= B)";
 | 
| 2893 |     16 | by (Blast_tac 1);
 | 
| 1531 |     17 | qed "subset_insert";
 | 
|  |     18 | 
 | 
| 923 |     19 | (*** Big Union -- least upper bound of a set  ***)
 | 
|  |     20 | 
 | 
| 5316 |     21 | Goal "B:A ==> B <= Union(A)";
 | 
|  |     22 | by (REPEAT (ares_tac [subsetI,UnionI] 1));
 | 
| 923 |     23 | qed "Union_upper";
 | 
|  |     24 | 
 | 
| 5316 |     25 | val [prem] = Goal "[| !!X. X:A ==> X<=C |] ==> Union(A) <= C";
 | 
| 1465 |     26 | by (rtac subsetI 1);
 | 
| 923 |     27 | by (REPEAT (eresolve_tac [asm_rl, UnionE, prem RS subsetD] 1));
 | 
|  |     28 | qed "Union_least";
 | 
|  |     29 | 
 | 
|  |     30 | (** General union **)
 | 
|  |     31 | 
 | 
| 5316 |     32 | Goal "a:A ==> B(a) <= (UN x:A. B(x))";
 | 
|  |     33 | by (Blast_tac 1);
 | 
| 923 |     34 | qed "UN_upper";
 | 
|  |     35 | 
 | 
| 5316 |     36 | val [prem] = Goal "[| !!x. x:A ==> B(x)<=C |] ==> (UN x:A. B(x)) <= C";
 | 
| 1465 |     37 | by (rtac subsetI 1);
 | 
| 923 |     38 | by (REPEAT (eresolve_tac [asm_rl, UN_E, prem RS subsetD] 1));
 | 
|  |     39 | qed "UN_least";
 | 
|  |     40 | 
 | 
|  |     41 | 
 | 
|  |     42 | (*** Big Intersection -- greatest lower bound of a set ***)
 | 
|  |     43 | 
 | 
| 5316 |     44 | Goal "B:A ==> Inter(A) <= B";
 | 
| 2893 |     45 | by (Blast_tac 1);
 | 
| 923 |     46 | qed "Inter_lower";
 | 
|  |     47 | 
 | 
| 5316 |     48 | val [prem] = Goal "[| !!X. X:A ==> C<=X |] ==> C <= Inter(A)";
 | 
| 1465 |     49 | by (rtac (InterI RS subsetI) 1);
 | 
| 923 |     50 | by (REPEAT (eresolve_tac [asm_rl, prem RS subsetD] 1));
 | 
|  |     51 | qed "Inter_greatest";
 | 
|  |     52 | 
 | 
| 5316 |     53 | Goal "a:A ==> (INT x:A. B(x)) <= B(a)";
 | 
|  |     54 | by (Blast_tac 1);
 | 
| 923 |     55 | qed "INT_lower";
 | 
|  |     56 | 
 | 
| 5316 |     57 | val [prem] = Goal "[| !!x. x:A ==> C<=B(x) |] ==> C <= (INT x:A. B(x))";
 | 
| 1465 |     58 | by (rtac (INT_I RS subsetI) 1);
 | 
| 923 |     59 | by (REPEAT (eresolve_tac [asm_rl, prem RS subsetD] 1));
 | 
|  |     60 | qed "INT_greatest";
 | 
|  |     61 | 
 | 
|  |     62 | (*** Finite Union -- the least upper bound of 2 sets ***)
 | 
|  |     63 | 
 | 
| 5316 |     64 | Goal "A <= A Un B";
 | 
| 2893 |     65 | by (Blast_tac 1);
 | 
| 923 |     66 | qed "Un_upper1";
 | 
|  |     67 | 
 | 
| 5316 |     68 | Goal "B <= A Un B";
 | 
| 2893 |     69 | by (Blast_tac 1);
 | 
| 923 |     70 | qed "Un_upper2";
 | 
|  |     71 | 
 | 
| 5316 |     72 | Goal "[| A<=C;  B<=C |] ==> A Un B <= C";
 | 
| 2893 |     73 | by (Blast_tac 1);
 | 
| 923 |     74 | qed "Un_least";
 | 
|  |     75 | 
 | 
|  |     76 | (*** Finite Intersection -- the greatest lower bound of 2 sets *)
 | 
|  |     77 | 
 | 
| 5316 |     78 | Goal "A Int B <= A";
 | 
| 2893 |     79 | by (Blast_tac 1);
 | 
| 923 |     80 | qed "Int_lower1";
 | 
|  |     81 | 
 | 
| 5316 |     82 | Goal "A Int B <= B";
 | 
| 2893 |     83 | by (Blast_tac 1);
 | 
| 923 |     84 | qed "Int_lower2";
 | 
|  |     85 | 
 | 
| 5316 |     86 | Goal "[| C<=A;  C<=B |] ==> C <= A Int B";
 | 
| 2893 |     87 | by (Blast_tac 1);
 | 
| 923 |     88 | qed "Int_greatest";
 | 
|  |     89 | 
 | 
|  |     90 | (*** Set difference ***)
 | 
|  |     91 | 
 | 
|  |     92 | qed_goal "Diff_subset" Set.thy "A-B <= (A::'a set)"
 | 
| 2893 |     93 |  (fn _ => [ (Blast_tac 1) ]);
 | 
| 923 |     94 | 
 | 
|  |     95 | (*** Monotonicity ***)
 | 
|  |     96 | 
 | 
| 5316 |     97 | Goal "mono(f) ==> f(A) Un f(B) <= f(A Un B)";
 | 
| 923 |     98 | by (rtac Un_least 1);
 | 
| 5316 |     99 | by (etac (Un_upper1 RSN (2,monoD)) 1);
 | 
|  |    100 | by (etac (Un_upper2 RSN (2,monoD)) 1);
 | 
| 923 |    101 | qed "mono_Un";
 | 
|  |    102 | 
 | 
| 5316 |    103 | Goal "mono(f) ==> f(A Int B) <= f(A) Int f(B)";
 | 
| 923 |    104 | by (rtac Int_greatest 1);
 | 
| 5316 |    105 | by (etac (Int_lower1 RSN (2,monoD)) 1);
 | 
|  |    106 | by (etac (Int_lower2 RSN (2,monoD)) 1);
 | 
| 923 |    107 | qed "mono_Int";
 |