| author | nipkow | 
| Mon, 23 Aug 1999 16:13:42 +0200 | |
| changeset 7323 | 16b7e2f1b4e3 | 
| parent 4347 | d683b7898c61 | 
| child 17456 | bcf7544875b2 | 
| permissions | -rw-r--r-- | 
| 1459 | 1 | (* Title: CCL/subset | 
| 0 | 2 | ID: $Id$ | 
| 3 | ||
| 4 | Modified version of | |
| 1459 | 5 | Title: HOL/subset | 
| 6 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 0 | 7 | Copyright 1991 University of Cambridge | 
| 8 | ||
| 9 | Derived rules involving subsets | |
| 10 | Union and Intersection as lattice operations | |
| 11 | *) | |
| 12 | ||
| 13 | (*** Big Union -- least upper bound of a set ***) | |
| 14 | ||
| 15 | val prems = goal Set.thy | |
| 16 | "B:A ==> B <= Union(A)"; | |
| 17 | by (REPEAT (ares_tac (prems@[subsetI,UnionI]) 1)); | |
| 757 | 18 | qed "Union_upper"; | 
| 0 | 19 | |
| 20 | val prems = goal Set.thy | |
| 21 | "[| !!X. X:A ==> X<=C |] ==> Union(A) <= C"; | |
| 22 | by (REPEAT (ares_tac [subsetI] 1 | |
| 23 | ORELSE eresolve_tac ([UnionE] @ (prems RL [subsetD])) 1)); | |
| 757 | 24 | qed "Union_least"; | 
| 0 | 25 | |
| 26 | ||
| 27 | (*** Big Intersection -- greatest lower bound of a set ***) | |
| 28 | ||
| 29 | val prems = goal Set.thy | |
| 30 | "B:A ==> Inter(A) <= B"; | |
| 31 | by (REPEAT (resolve_tac (prems@[subsetI]) 1 | |
| 32 | ORELSE etac InterD 1)); | |
| 757 | 33 | qed "Inter_lower"; | 
| 0 | 34 | |
| 35 | val prems = goal Set.thy | |
| 36 | "[| !!X. X:A ==> C<=X |] ==> C <= Inter(A)"; | |
| 37 | by (REPEAT (ares_tac [subsetI,InterI] 1 | |
| 38 | ORELSE eresolve_tac (prems RL [subsetD]) 1)); | |
| 757 | 39 | qed "Inter_greatest"; | 
| 0 | 40 | |
| 41 | (*** Finite Union -- the least upper bound of 2 sets ***) | |
| 42 | ||
| 43 | goal Set.thy "A <= A Un B"; | |
| 44 | by (REPEAT (ares_tac [subsetI,UnI1] 1)); | |
| 757 | 45 | qed "Un_upper1"; | 
| 0 | 46 | |
| 47 | goal Set.thy "B <= A Un B"; | |
| 48 | by (REPEAT (ares_tac [subsetI,UnI2] 1)); | |
| 757 | 49 | qed "Un_upper2"; | 
| 0 | 50 | |
| 51 | val prems = goal Set.thy "[| A<=C; B<=C |] ==> A Un B <= C"; | |
| 52 | by (cut_facts_tac prems 1); | |
| 53 | by (DEPTH_SOLVE (ares_tac [subsetI] 1 | |
| 54 | ORELSE eresolve_tac [UnE,subsetD] 1)); | |
| 757 | 55 | qed "Un_least"; | 
| 0 | 56 | |
| 57 | (*** Finite Intersection -- the greatest lower bound of 2 sets *) | |
| 58 | ||
| 59 | goal Set.thy "A Int B <= A"; | |
| 60 | by (REPEAT (ares_tac [subsetI] 1 ORELSE etac IntE 1)); | |
| 757 | 61 | qed "Int_lower1"; | 
| 0 | 62 | |
| 63 | goal Set.thy "A Int B <= B"; | |
| 64 | by (REPEAT (ares_tac [subsetI] 1 ORELSE etac IntE 1)); | |
| 757 | 65 | qed "Int_lower2"; | 
| 0 | 66 | |
| 67 | val prems = goal Set.thy "[| C<=A; C<=B |] ==> C <= A Int B"; | |
| 68 | by (cut_facts_tac prems 1); | |
| 69 | by (REPEAT (ares_tac [subsetI,IntI] 1 | |
| 70 | ORELSE etac subsetD 1)); | |
| 757 | 71 | qed "Int_greatest"; | 
| 0 | 72 | |
| 73 | (*** Monotonicity ***) | |
| 74 | ||
| 75 | val [prem] = goalw Set.thy [mono_def] | |
| 76 | "[| !!A B. A <= B ==> f(A) <= f(B) |] ==> mono(f)"; | |
| 77 | by (REPEAT (ares_tac [allI, impI, prem] 1)); | |
| 757 | 78 | qed "monoI"; | 
| 0 | 79 | |
| 80 | val [major,minor] = goalw Set.thy [mono_def] | |
| 81 | "[| mono(f); A <= B |] ==> f(A) <= f(B)"; | |
| 82 | by (rtac (major RS spec RS spec RS mp) 1); | |
| 83 | by (rtac minor 1); | |
| 757 | 84 | qed "monoD"; | 
| 0 | 85 | |
| 86 | val [prem] = goal Set.thy "mono(f) ==> f(A) Un f(B) <= f(A Un B)"; | |
| 87 | by (rtac Un_least 1); | |
| 88 | by (rtac (Un_upper1 RS (prem RS monoD)) 1); | |
| 89 | by (rtac (Un_upper2 RS (prem RS monoD)) 1); | |
| 757 | 90 | qed "mono_Un"; | 
| 0 | 91 | |
| 92 | val [prem] = goal Set.thy "mono(f) ==> f(A Int B) <= f(A) Int f(B)"; | |
| 93 | by (rtac Int_greatest 1); | |
| 94 | by (rtac (Int_lower1 RS (prem RS monoD)) 1); | |
| 95 | by (rtac (Int_lower2 RS (prem RS monoD)) 1); | |
| 757 | 96 | qed "mono_Int"; | 
| 0 | 97 | |
| 98 | (****) | |
| 99 | ||
| 100 | val set_cs = FOL_cs | |
| 101 | addSIs [ballI, subsetI, InterI, INT_I, CollectI, | |
| 1459 | 102 | ComplI, IntI, UnCI, singletonI] | 
| 0 | 103 | addIs [bexI, UnionI, UN_I] | 
| 104 | addSEs [bexE, UnionE, UN_E, | |
| 1459 | 105 | CollectE, ComplE, IntE, UnE, emptyE, singletonE] | 
| 0 | 106 | addEs [ballE, InterD, InterE, INT_D, INT_E, subsetD, subsetCE]; | 
| 107 | ||
| 108 | fun cfast_tac prems = cut_facts_tac prems THEN' fast_tac set_cs; | |
| 109 | ||
| 110 | fun prover s = prove_goal Set.thy s (fn _=>[fast_tac set_cs 1]); | |
| 111 | ||
| 112 | val mem_rews = [trivial_set,empty_eq] @ (map prover | |
| 113 | [ "(a : A Un B) <-> (a:A | a:B)", | |
| 114 | "(a : A Int B) <-> (a:A & a:B)", | |
| 115 | "(a : Compl(B)) <-> (~a:B)", | |
| 116 |    "(a : {b})      <->  (a=b)",
 | |
| 117 |    "(a : {})       <->   False",
 | |
| 3837 | 118 |    "(a : {x. P(x)}) <->  P(a)" ]);
 | 
| 0 | 119 | |
| 8 
c3d2c6dcf3f0
Installation of new simplfier.  Previously appeared to set up the old
 lcp parents: 
0diff
changeset | 120 | val set_congs = [ball_cong, bex_cong, INT_cong, UN_cong]; | 
| 0 | 121 | |
| 1963 | 122 | val set_ss = FOL_ss addcongs set_congs | 
| 2035 | 123 | addsimps mem_rews; |