| 
17456
 | 
     1  | 
(*  Title:      CCL/Set.thy
  | 
| 
0
 | 
     2  | 
    ID:         $Id$
  | 
| 
 | 
     3  | 
*)
  | 
| 
 | 
     4  | 
  | 
| 
17456
 | 
     5  | 
header {* Extending FOL by a modified version of HOL set theory *}
 | 
| 
 | 
     6  | 
  | 
| 
 | 
     7  | 
theory Set
  | 
| 
 | 
     8  | 
imports FOL
  | 
| 
 | 
     9  | 
begin
  | 
| 
0
 | 
    10  | 
  | 
| 
3935
 | 
    11  | 
global
  | 
| 
 | 
    12  | 
  | 
| 
17456
 | 
    13  | 
typedecl 'a set
  | 
| 
 | 
    14  | 
arities set :: ("term") "term"
 | 
| 
0
 | 
    15  | 
  | 
| 
 | 
    16  | 
consts
  | 
| 
 | 
    17  | 
  Collect       :: "['a => o] => 'a set"                    (*comprehension*)
  | 
| 
 | 
    18  | 
  Compl         :: "('a set) => 'a set"                     (*complement*)
 | 
| 
24825
 | 
    19  | 
  Int           :: "['a set, 'a set] => 'a set"         (infixl "Int" 70)
  | 
| 
 | 
    20  | 
  Un            :: "['a set, 'a set] => 'a set"         (infixl "Un" 65)
  | 
| 
17456
 | 
    21  | 
  Union         :: "(('a set)set) => 'a set"                (*...of a set*)
 | 
| 
 | 
    22  | 
  Inter         :: "(('a set)set) => 'a set"                (*...of a set*)
 | 
| 
 | 
    23  | 
  UNION         :: "['a set, 'a => 'b set] => 'b set"       (*general*)
  | 
| 
 | 
    24  | 
  INTER         :: "['a set, 'a => 'b set] => 'b set"       (*general*)
  | 
| 
 | 
    25  | 
  Ball          :: "['a set, 'a => o] => o"                 (*bounded quants*)
  | 
| 
 | 
    26  | 
  Bex           :: "['a set, 'a => o] => o"                 (*bounded quants*)
  | 
| 
0
 | 
    27  | 
  mono          :: "['a set => 'b set] => o"                (*monotonicity*)
  | 
| 
24825
 | 
    28  | 
  mem           :: "['a, 'a set] => o"                  (infixl ":" 50) (*membership*)
  | 
| 
 | 
    29  | 
  subset        :: "['a set, 'a set] => o"              (infixl "<=" 50)
  | 
| 
0
 | 
    30  | 
  singleton     :: "'a => 'a set"                       ("{_}")
 | 
| 
 | 
    31  | 
  empty         :: "'a set"                             ("{}")
 | 
| 
 | 
    32  | 
  | 
| 
3935
 | 
    33  | 
syntax
  | 
| 
0
 | 
    34  | 
  "@Coll"       :: "[idt, o] => 'a set"                 ("(1{_./ _})") (*collection*)
 | 
| 
 | 
    35  | 
  | 
| 
 | 
    36  | 
  (* Big Intersection / Union *)
  | 
| 
 | 
    37  | 
  | 
| 
 | 
    38  | 
  "@INTER"      :: "[idt, 'a set, 'b set] => 'b set"    ("(INT _:_./ _)" [0, 0, 0] 10)
 | 
| 
 | 
    39  | 
  "@UNION"      :: "[idt, 'a set, 'b set] => 'b set"    ("(UN _:_./ _)" [0, 0, 0] 10)
 | 
| 
 | 
    40  | 
  | 
| 
 | 
    41  | 
  (* Bounded Quantifiers *)
  | 
| 
 | 
    42  | 
  | 
| 
 | 
    43  | 
  "@Ball"       :: "[idt, 'a set, o] => o"              ("(ALL _:_./ _)" [0, 0, 0] 10)
 | 
| 
 | 
    44  | 
  "@Bex"        :: "[idt, 'a set, o] => o"              ("(EX _:_./ _)" [0, 0, 0] 10)
 | 
| 
 | 
    45  | 
  | 
| 
 | 
    46  | 
translations
  | 
| 
 | 
    47  | 
  "{x. P}"      == "Collect(%x. P)"
 | 
| 
 | 
    48  | 
  "INT x:A. B"  == "INTER(A, %x. B)"
  | 
| 
 | 
    49  | 
  "UN x:A. B"   == "UNION(A, %x. B)"
  | 
| 
 | 
    50  | 
  "ALL x:A. P"  == "Ball(A, %x. P)"
  | 
| 
 | 
    51  | 
  "EX x:A. P"   == "Bex(A, %x. P)"
  | 
| 
 | 
    52  | 
  | 
| 
3935
 | 
    53  | 
local
  | 
| 
0
 | 
    54  | 
  | 
| 
17456
 | 
    55  | 
axioms
  | 
| 
 | 
    56  | 
  mem_Collect_iff:       "(a : {x. P(x)}) <-> P(a)"
 | 
| 
 | 
    57  | 
  set_extension:         "A=B <-> (ALL x. x:A <-> x:B)"
  | 
| 
0
 | 
    58  | 
  | 
| 
17456
 | 
    59  | 
defs
  | 
| 
 | 
    60  | 
  Ball_def:      "Ball(A, P)  == ALL x. x:A --> P(x)"
  | 
| 
 | 
    61  | 
  Bex_def:       "Bex(A, P)   == EX x. x:A & P(x)"
  | 
| 
 | 
    62  | 
  mono_def:      "mono(f)     == (ALL A B. A <= B --> f(A) <= f(B))"
  | 
| 
 | 
    63  | 
  subset_def:    "A <= B      == ALL x:A. x:B"
  | 
| 
 | 
    64  | 
  singleton_def: "{a}         == {x. x=a}"
 | 
| 
 | 
    65  | 
  empty_def:     "{}          == {x. False}"
 | 
| 
 | 
    66  | 
  Un_def:        "A Un B      == {x. x:A | x:B}"
 | 
| 
 | 
    67  | 
  Int_def:       "A Int B     == {x. x:A & x:B}"
 | 
| 
 | 
    68  | 
  Compl_def:     "Compl(A)    == {x. ~x:A}"
 | 
| 
 | 
    69  | 
  INTER_def:     "INTER(A, B) == {y. ALL x:A. y: B(x)}"
 | 
| 
 | 
    70  | 
  UNION_def:     "UNION(A, B) == {y. EX x:A. y: B(x)}"
 | 
| 
 | 
    71  | 
  Inter_def:     "Inter(S)    == (INT x:S. x)"
  | 
| 
 | 
    72  | 
  Union_def:     "Union(S)    == (UN x:S. x)"
  | 
| 
 | 
    73  | 
  | 
| 
20140
 | 
    74  | 
  | 
| 
 | 
    75  | 
lemma CollectI: "[| P(a) |] ==> a : {x. P(x)}"
 | 
| 
 | 
    76  | 
  apply (rule mem_Collect_iff [THEN iffD2])
  | 
| 
 | 
    77  | 
  apply assumption
  | 
| 
 | 
    78  | 
  done
  | 
| 
 | 
    79  | 
  | 
| 
 | 
    80  | 
lemma CollectD: "[| a : {x. P(x)} |] ==> P(a)"
 | 
| 
 | 
    81  | 
  apply (erule mem_Collect_iff [THEN iffD1])
  | 
| 
 | 
    82  | 
  done
  | 
| 
 | 
    83  | 
  | 
| 
 | 
    84  | 
lemmas CollectE = CollectD [elim_format]
  | 
| 
 | 
    85  | 
  | 
| 
 | 
    86  | 
lemma set_ext: "[| !!x. x:A <-> x:B |] ==> A = B"
  | 
| 
 | 
    87  | 
  apply (rule set_extension [THEN iffD2])
  | 
| 
 | 
    88  | 
  apply simp
  | 
| 
 | 
    89  | 
  done
  | 
| 
 | 
    90  | 
  | 
| 
 | 
    91  | 
  | 
| 
 | 
    92  | 
subsection {* Bounded quantifiers *}
 | 
| 
 | 
    93  | 
  | 
| 
 | 
    94  | 
lemma ballI: "[| !!x. x:A ==> P(x) |] ==> ALL x:A. P(x)"
  | 
| 
 | 
    95  | 
  by (simp add: Ball_def)
  | 
| 
 | 
    96  | 
  | 
| 
 | 
    97  | 
lemma bspec: "[| ALL x:A. P(x);  x:A |] ==> P(x)"
  | 
| 
 | 
    98  | 
  by (simp add: Ball_def)
  | 
| 
 | 
    99  | 
  | 
| 
 | 
   100  | 
lemma ballE: "[| ALL x:A. P(x);  P(x) ==> Q;  ~ x:A ==> Q |] ==> Q"
  | 
| 
 | 
   101  | 
  unfolding Ball_def by blast
  | 
| 
 | 
   102  | 
  | 
| 
 | 
   103  | 
lemma bexI: "[| P(x);  x:A |] ==> EX x:A. P(x)"
  | 
| 
 | 
   104  | 
  unfolding Bex_def by blast
  | 
| 
 | 
   105  | 
  | 
| 
 | 
   106  | 
lemma bexCI: "[| EX x:A. ~P(x) ==> P(a);  a:A |] ==> EX x:A. P(x)"
  | 
| 
 | 
   107  | 
  unfolding Bex_def by blast
  | 
| 
 | 
   108  | 
  | 
| 
 | 
   109  | 
lemma bexE: "[| EX x:A. P(x);  !!x. [| x:A; P(x) |] ==> Q  |] ==> Q"
  | 
| 
 | 
   110  | 
  unfolding Bex_def by blast
  | 
| 
 | 
   111  | 
  | 
| 
 | 
   112  | 
(*Trival rewrite rule;   (! x:A.P)=P holds only if A is nonempty!*)
  | 
| 
 | 
   113  | 
lemma ball_rew: "(ALL x:A. True) <-> True"
  | 
| 
 | 
   114  | 
  by (blast intro: ballI)
  | 
| 
 | 
   115  | 
  | 
| 
 | 
   116  | 
  | 
| 
 | 
   117  | 
subsection {* Congruence rules *}
 | 
| 
 | 
   118  | 
  | 
| 
 | 
   119  | 
lemma ball_cong:
  | 
| 
 | 
   120  | 
  "[| A=A';  !!x. x:A' ==> P(x) <-> P'(x) |] ==>
  | 
| 
 | 
   121  | 
    (ALL x:A. P(x)) <-> (ALL x:A'. P'(x))"
  | 
| 
 | 
   122  | 
  by (blast intro: ballI elim: ballE)
  | 
| 
 | 
   123  | 
  | 
| 
 | 
   124  | 
lemma bex_cong:
  | 
| 
 | 
   125  | 
  "[| A=A';  !!x. x:A' ==> P(x) <-> P'(x) |] ==>
  | 
| 
 | 
   126  | 
    (EX x:A. P(x)) <-> (EX x:A'. P'(x))"
  | 
| 
 | 
   127  | 
  by (blast intro: bexI elim: bexE)
  | 
| 
 | 
   128  | 
  | 
| 
 | 
   129  | 
  | 
| 
 | 
   130  | 
subsection {* Rules for subsets *}
 | 
| 
 | 
   131  | 
  | 
| 
 | 
   132  | 
lemma subsetI: "(!!x. x:A ==> x:B) ==> A <= B"
  | 
| 
 | 
   133  | 
  unfolding subset_def by (blast intro: ballI)
  | 
| 
 | 
   134  | 
  | 
| 
 | 
   135  | 
(*Rule in Modus Ponens style*)
  | 
| 
 | 
   136  | 
lemma subsetD: "[| A <= B;  c:A |] ==> c:B"
  | 
| 
 | 
   137  | 
  unfolding subset_def by (blast elim: ballE)
  | 
| 
 | 
   138  | 
  | 
| 
 | 
   139  | 
(*Classical elimination rule*)
  | 
| 
 | 
   140  | 
lemma subsetCE: "[| A <= B;  ~(c:A) ==> P;  c:B ==> P |] ==> P"
  | 
| 
 | 
   141  | 
  by (blast dest: subsetD)
  | 
| 
 | 
   142  | 
  | 
| 
 | 
   143  | 
lemma subset_refl: "A <= A"
  | 
| 
 | 
   144  | 
  by (blast intro: subsetI)
  | 
| 
 | 
   145  | 
  | 
| 
 | 
   146  | 
lemma subset_trans: "[| A<=B;  B<=C |] ==> A<=C"
  | 
| 
 | 
   147  | 
  by (blast intro: subsetI dest: subsetD)
  | 
| 
 | 
   148  | 
  | 
| 
 | 
   149  | 
  | 
| 
 | 
   150  | 
subsection {* Rules for equality *}
 | 
| 
 | 
   151  | 
  | 
| 
 | 
   152  | 
(*Anti-symmetry of the subset relation*)
  | 
| 
 | 
   153  | 
lemma subset_antisym: "[| A <= B;  B <= A |] ==> A = B"
  | 
| 
 | 
   154  | 
  by (blast intro: set_ext dest: subsetD)
  | 
| 
 | 
   155  | 
  | 
| 
 | 
   156  | 
lemmas equalityI = subset_antisym
  | 
| 
 | 
   157  | 
  | 
| 
 | 
   158  | 
(* Equality rules from ZF set theory -- are they appropriate here? *)
  | 
| 
 | 
   159  | 
lemma equalityD1: "A = B ==> A<=B"
  | 
| 
 | 
   160  | 
  and equalityD2: "A = B ==> B<=A"
  | 
| 
 | 
   161  | 
  by (simp_all add: subset_refl)
  | 
| 
 | 
   162  | 
  | 
| 
 | 
   163  | 
lemma equalityE: "[| A = B;  [| A<=B; B<=A |] ==> P |]  ==>  P"
  | 
| 
 | 
   164  | 
  by (simp add: subset_refl)
  | 
| 
 | 
   165  | 
  | 
| 
 | 
   166  | 
lemma equalityCE:
  | 
| 
 | 
   167  | 
    "[| A = B;  [| c:A; c:B |] ==> P;  [| ~ c:A; ~ c:B |] ==> P |]  ==>  P"
  | 
| 
 | 
   168  | 
  by (blast elim: equalityE subsetCE)
  | 
| 
 | 
   169  | 
  | 
| 
 | 
   170  | 
lemma trivial_set: "{x. x:A} = A"
 | 
| 
 | 
   171  | 
  by (blast intro: equalityI subsetI CollectI dest: CollectD)
  | 
| 
 | 
   172  | 
  | 
| 
 | 
   173  | 
  | 
| 
 | 
   174  | 
subsection {* Rules for binary union *}
 | 
| 
 | 
   175  | 
  | 
| 
 | 
   176  | 
lemma UnI1: "c:A ==> c : A Un B"
  | 
| 
 | 
   177  | 
  and UnI2: "c:B ==> c : A Un B"
  | 
| 
 | 
   178  | 
  unfolding Un_def by (blast intro: CollectI)+
  | 
| 
 | 
   179  | 
  | 
| 
 | 
   180  | 
(*Classical introduction rule: no commitment to A vs B*)
  | 
| 
 | 
   181  | 
lemma UnCI: "(~c:B ==> c:A) ==> c : A Un B"
  | 
| 
 | 
   182  | 
  by (blast intro: UnI1 UnI2)
  | 
| 
 | 
   183  | 
  | 
| 
 | 
   184  | 
lemma UnE: "[| c : A Un B;  c:A ==> P;  c:B ==> P |] ==> P"
  | 
| 
 | 
   185  | 
  unfolding Un_def by (blast dest: CollectD)
  | 
| 
 | 
   186  | 
  | 
| 
 | 
   187  | 
  | 
| 
 | 
   188  | 
subsection {* Rules for small intersection *}
 | 
| 
 | 
   189  | 
  | 
| 
 | 
   190  | 
lemma IntI: "[| c:A;  c:B |] ==> c : A Int B"
  | 
| 
 | 
   191  | 
  unfolding Int_def by (blast intro: CollectI)
  | 
| 
 | 
   192  | 
  | 
| 
 | 
   193  | 
lemma IntD1: "c : A Int B ==> c:A"
  | 
| 
 | 
   194  | 
  and IntD2: "c : A Int B ==> c:B"
  | 
| 
 | 
   195  | 
  unfolding Int_def by (blast dest: CollectD)+
  | 
| 
 | 
   196  | 
  | 
| 
 | 
   197  | 
lemma IntE: "[| c : A Int B;  [| c:A; c:B |] ==> P |] ==> P"
  | 
| 
 | 
   198  | 
  by (blast dest: IntD1 IntD2)
  | 
| 
 | 
   199  | 
  | 
| 
 | 
   200  | 
  | 
| 
 | 
   201  | 
subsection {* Rules for set complement *}
 | 
| 
 | 
   202  | 
  | 
| 
 | 
   203  | 
lemma ComplI: "[| c:A ==> False |] ==> c : Compl(A)"
  | 
| 
 | 
   204  | 
  unfolding Compl_def by (blast intro: CollectI)
  | 
| 
 | 
   205  | 
  | 
| 
 | 
   206  | 
(*This form, with negated conclusion, works well with the Classical prover.
  | 
| 
 | 
   207  | 
  Negated assumptions behave like formulae on the right side of the notional
  | 
| 
 | 
   208  | 
  turnstile...*)
  | 
| 
 | 
   209  | 
lemma ComplD: "[| c : Compl(A) |] ==> ~c:A"
  | 
| 
 | 
   210  | 
  unfolding Compl_def by (blast dest: CollectD)
  | 
| 
 | 
   211  | 
  | 
| 
 | 
   212  | 
lemmas ComplE = ComplD [elim_format]
  | 
| 
 | 
   213  | 
  | 
| 
 | 
   214  | 
  | 
| 
 | 
   215  | 
subsection {* Empty sets *}
 | 
| 
 | 
   216  | 
  | 
| 
 | 
   217  | 
lemma empty_eq: "{x. False} = {}"
 | 
| 
 | 
   218  | 
  by (simp add: empty_def)
  | 
| 
 | 
   219  | 
  | 
| 
 | 
   220  | 
lemma emptyD: "a : {} ==> P"
 | 
| 
 | 
   221  | 
  unfolding empty_def by (blast dest: CollectD)
  | 
| 
 | 
   222  | 
  | 
| 
 | 
   223  | 
lemmas emptyE = emptyD [elim_format]
  | 
| 
 | 
   224  | 
  | 
| 
 | 
   225  | 
lemma not_emptyD:
  | 
| 
 | 
   226  | 
  assumes "~ A={}"
 | 
| 
 | 
   227  | 
  shows "EX x. x:A"
  | 
| 
 | 
   228  | 
proof -
  | 
| 
 | 
   229  | 
  have "\<not> (EX x. x:A) \<Longrightarrow> A = {}"
 | 
| 
 | 
   230  | 
    by (rule equalityI) (blast intro!: subsetI elim!: emptyD)+
  | 
| 
 | 
   231  | 
  with prems show ?thesis by blast
  | 
| 
 | 
   232  | 
qed
  | 
| 
 | 
   233  | 
  | 
| 
 | 
   234  | 
  | 
| 
 | 
   235  | 
subsection {* Singleton sets *}
 | 
| 
 | 
   236  | 
  | 
| 
 | 
   237  | 
lemma singletonI: "a : {a}"
 | 
| 
 | 
   238  | 
  unfolding singleton_def by (blast intro: CollectI)
  | 
| 
 | 
   239  | 
  | 
| 
 | 
   240  | 
lemma singletonD: "b : {a} ==> b=a"
 | 
| 
 | 
   241  | 
  unfolding singleton_def by (blast dest: CollectD)
  | 
| 
 | 
   242  | 
  | 
| 
 | 
   243  | 
lemmas singletonE = singletonD [elim_format]
  | 
| 
 | 
   244  | 
  | 
| 
 | 
   245  | 
  | 
| 
 | 
   246  | 
subsection {* Unions of families *}
 | 
| 
 | 
   247  | 
  | 
| 
 | 
   248  | 
(*The order of the premises presupposes that A is rigid; b may be flexible*)
  | 
| 
 | 
   249  | 
lemma UN_I: "[| a:A;  b: B(a) |] ==> b: (UN x:A. B(x))"
  | 
| 
 | 
   250  | 
  unfolding UNION_def by (blast intro: bexI CollectI)
  | 
| 
 | 
   251  | 
  | 
| 
 | 
   252  | 
lemma UN_E: "[| b : (UN x:A. B(x));  !!x.[| x:A;  b: B(x) |] ==> R |] ==> R"
  | 
| 
 | 
   253  | 
  unfolding UNION_def by (blast dest: CollectD elim: bexE)
  | 
| 
 | 
   254  | 
  | 
| 
 | 
   255  | 
lemma UN_cong:
  | 
| 
 | 
   256  | 
  "[| A=B;  !!x. x:B ==> C(x) = D(x) |] ==>
  | 
| 
 | 
   257  | 
    (UN x:A. C(x)) = (UN x:B. D(x))"
  | 
| 
 | 
   258  | 
  by (simp add: UNION_def cong: bex_cong)
  | 
| 
 | 
   259  | 
  | 
| 
 | 
   260  | 
  | 
| 
 | 
   261  | 
subsection {* Intersections of families *}
 | 
| 
 | 
   262  | 
  | 
| 
 | 
   263  | 
lemma INT_I: "(!!x. x:A ==> b: B(x)) ==> b : (INT x:A. B(x))"
  | 
| 
 | 
   264  | 
  unfolding INTER_def by (blast intro: CollectI ballI)
  | 
| 
 | 
   265  | 
  | 
| 
 | 
   266  | 
lemma INT_D: "[| b : (INT x:A. B(x));  a:A |] ==> b: B(a)"
  | 
| 
 | 
   267  | 
  unfolding INTER_def by (blast dest: CollectD bspec)
  | 
| 
 | 
   268  | 
  | 
| 
 | 
   269  | 
(*"Classical" elimination rule -- does not require proving X:C *)
  | 
| 
 | 
   270  | 
lemma INT_E: "[| b : (INT x:A. B(x));  b: B(a) ==> R;  ~ a:A ==> R |] ==> R"
  | 
| 
 | 
   271  | 
  unfolding INTER_def by (blast dest: CollectD bspec)
  | 
| 
 | 
   272  | 
  | 
| 
 | 
   273  | 
lemma INT_cong:
  | 
| 
 | 
   274  | 
  "[| A=B;  !!x. x:B ==> C(x) = D(x) |] ==>
  | 
| 
 | 
   275  | 
    (INT x:A. C(x)) = (INT x:B. D(x))"
  | 
| 
 | 
   276  | 
  by (simp add: INTER_def cong: ball_cong)
  | 
| 
 | 
   277  | 
  | 
| 
 | 
   278  | 
  | 
| 
 | 
   279  | 
subsection {* Rules for Unions *}
 | 
| 
 | 
   280  | 
  | 
| 
 | 
   281  | 
(*The order of the premises presupposes that C is rigid; A may be flexible*)
  | 
| 
 | 
   282  | 
lemma UnionI: "[| X:C;  A:X |] ==> A : Union(C)"
  | 
| 
 | 
   283  | 
  unfolding Union_def by (blast intro: UN_I)
  | 
| 
 | 
   284  | 
  | 
| 
 | 
   285  | 
lemma UnionE: "[| A : Union(C);  !!X.[| A:X;  X:C |] ==> R |] ==> R"
  | 
| 
 | 
   286  | 
  unfolding Union_def by (blast elim: UN_E)
  | 
| 
 | 
   287  | 
  | 
| 
 | 
   288  | 
  | 
| 
 | 
   289  | 
subsection {* Rules for Inter *}
 | 
| 
 | 
   290  | 
  | 
| 
 | 
   291  | 
lemma InterI: "[| !!X. X:C ==> A:X |] ==> A : Inter(C)"
  | 
| 
 | 
   292  | 
  unfolding Inter_def by (blast intro: INT_I)
  | 
| 
 | 
   293  | 
  | 
| 
 | 
   294  | 
(*A "destruct" rule -- every X in C contains A as an element, but
  | 
| 
 | 
   295  | 
  A:X can hold when X:C does not!  This rule is analogous to "spec". *)
  | 
| 
 | 
   296  | 
lemma InterD: "[| A : Inter(C);  X:C |] ==> A:X"
  | 
| 
 | 
   297  | 
  unfolding Inter_def by (blast dest: INT_D)
  | 
| 
 | 
   298  | 
  | 
| 
 | 
   299  | 
(*"Classical" elimination rule -- does not require proving X:C *)
  | 
| 
 | 
   300  | 
lemma InterE: "[| A : Inter(C);  A:X ==> R;  ~ X:C ==> R |] ==> R"
  | 
| 
 | 
   301  | 
  unfolding Inter_def by (blast elim: INT_E)
  | 
| 
 | 
   302  | 
  | 
| 
 | 
   303  | 
  | 
| 
 | 
   304  | 
section {* Derived rules involving subsets; Union and Intersection as lattice operations *}
 | 
| 
 | 
   305  | 
  | 
| 
 | 
   306  | 
subsection {* Big Union -- least upper bound of a set *}
 | 
| 
 | 
   307  | 
  | 
| 
 | 
   308  | 
lemma Union_upper: "B:A ==> B <= Union(A)"
  | 
| 
 | 
   309  | 
  by (blast intro: subsetI UnionI)
  | 
| 
 | 
   310  | 
  | 
| 
 | 
   311  | 
lemma Union_least: "[| !!X. X:A ==> X<=C |] ==> Union(A) <= C"
  | 
| 
 | 
   312  | 
  by (blast intro: subsetI dest: subsetD elim: UnionE)
  | 
| 
 | 
   313  | 
  | 
| 
 | 
   314  | 
  | 
| 
 | 
   315  | 
subsection {* Big Intersection -- greatest lower bound of a set *}
 | 
| 
 | 
   316  | 
  | 
| 
 | 
   317  | 
lemma Inter_lower: "B:A ==> Inter(A) <= B"
  | 
| 
 | 
   318  | 
  by (blast intro: subsetI dest: InterD)
  | 
| 
 | 
   319  | 
  | 
| 
 | 
   320  | 
lemma Inter_greatest: "[| !!X. X:A ==> C<=X |] ==> C <= Inter(A)"
  | 
| 
 | 
   321  | 
  by (blast intro: subsetI InterI dest: subsetD)
  | 
| 
 | 
   322  | 
  | 
| 
 | 
   323  | 
  | 
| 
 | 
   324  | 
subsection {* Finite Union -- the least upper bound of 2 sets *}
 | 
| 
 | 
   325  | 
  | 
| 
 | 
   326  | 
lemma Un_upper1: "A <= A Un B"
  | 
| 
 | 
   327  | 
  by (blast intro: subsetI UnI1)
  | 
| 
 | 
   328  | 
  | 
| 
 | 
   329  | 
lemma Un_upper2: "B <= A Un B"
  | 
| 
 | 
   330  | 
  by (blast intro: subsetI UnI2)
  | 
| 
 | 
   331  | 
  | 
| 
 | 
   332  | 
lemma Un_least: "[| A<=C;  B<=C |] ==> A Un B <= C"
  | 
| 
 | 
   333  | 
  by (blast intro: subsetI elim: UnE dest: subsetD)
  | 
| 
 | 
   334  | 
  | 
| 
 | 
   335  | 
  | 
| 
 | 
   336  | 
subsection {* Finite Intersection -- the greatest lower bound of 2 sets *}
 | 
| 
 | 
   337  | 
  | 
| 
 | 
   338  | 
lemma Int_lower1: "A Int B <= A"
  | 
| 
 | 
   339  | 
  by (blast intro: subsetI elim: IntE)
  | 
| 
 | 
   340  | 
  | 
| 
 | 
   341  | 
lemma Int_lower2: "A Int B <= B"
  | 
| 
 | 
   342  | 
  by (blast intro: subsetI elim: IntE)
  | 
| 
 | 
   343  | 
  | 
| 
 | 
   344  | 
lemma Int_greatest: "[| C<=A;  C<=B |] ==> C <= A Int B"
  | 
| 
 | 
   345  | 
  by (blast intro: subsetI IntI dest: subsetD)
  | 
| 
 | 
   346  | 
  | 
| 
 | 
   347  | 
  | 
| 
 | 
   348  | 
subsection {* Monotonicity *}
 | 
| 
 | 
   349  | 
  | 
| 
 | 
   350  | 
lemma monoI: "[| !!A B. A <= B ==> f(A) <= f(B) |] ==> mono(f)"
  | 
| 
 | 
   351  | 
  unfolding mono_def by blast
  | 
| 
 | 
   352  | 
  | 
| 
 | 
   353  | 
lemma monoD: "[| mono(f);  A <= B |] ==> f(A) <= f(B)"
  | 
| 
 | 
   354  | 
  unfolding mono_def by blast
  | 
| 
 | 
   355  | 
  | 
| 
 | 
   356  | 
lemma mono_Un: "mono(f) ==> f(A) Un f(B) <= f(A Un B)"
  | 
| 
 | 
   357  | 
  by (blast intro: Un_least dest: monoD intro: Un_upper1 Un_upper2)
  | 
| 
 | 
   358  | 
  | 
| 
 | 
   359  | 
lemma mono_Int: "mono(f) ==> f(A Int B) <= f(A) Int f(B)"
  | 
| 
 | 
   360  | 
  by (blast intro: Int_greatest dest: monoD intro: Int_lower1 Int_lower2)
  | 
| 
 | 
   361  | 
  | 
| 
 | 
   362  | 
  | 
| 
 | 
   363  | 
subsection {* Automated reasoning setup *}
 | 
| 
 | 
   364  | 
  | 
| 
 | 
   365  | 
lemmas [intro!] = ballI subsetI InterI INT_I CollectI ComplI IntI UnCI singletonI
  | 
| 
 | 
   366  | 
  and [intro] = bexI UnionI UN_I
  | 
| 
 | 
   367  | 
  and [elim!] = bexE UnionE UN_E CollectE ComplE IntE UnE emptyE singletonE
  | 
| 
 | 
   368  | 
  and [elim] = ballE InterD InterE INT_D INT_E subsetD subsetCE
  | 
| 
 | 
   369  | 
  | 
| 
 | 
   370  | 
lemma mem_rews:
  | 
| 
 | 
   371  | 
  "(a : A Un B)   <->  (a:A | a:B)"
  | 
| 
 | 
   372  | 
  "(a : A Int B)  <->  (a:A & a:B)"
  | 
| 
 | 
   373  | 
  "(a : Compl(B)) <->  (~a:B)"
  | 
| 
 | 
   374  | 
  "(a : {b})      <->  (a=b)"
 | 
| 
 | 
   375  | 
  "(a : {})       <->   False"
 | 
| 
 | 
   376  | 
  "(a : {x. P(x)}) <->  P(a)"
 | 
| 
 | 
   377  | 
  by blast+
  | 
| 
 | 
   378  | 
  | 
| 
 | 
   379  | 
lemmas [simp] = trivial_set empty_eq mem_rews
  | 
| 
 | 
   380  | 
  and [cong] = ball_cong bex_cong INT_cong UN_cong
  | 
| 
 | 
   381  | 
  | 
| 
 | 
   382  | 
  | 
| 
 | 
   383  | 
section {* Equalities involving union, intersection, inclusion, etc. *}
 | 
| 
 | 
   384  | 
  | 
| 
 | 
   385  | 
subsection {* Binary Intersection *}
 | 
| 
 | 
   386  | 
  | 
| 
 | 
   387  | 
lemma Int_absorb: "A Int A = A"
  | 
| 
 | 
   388  | 
  by (blast intro: equalityI)
  | 
| 
 | 
   389  | 
  | 
| 
 | 
   390  | 
lemma Int_commute: "A Int B  =  B Int A"
  | 
| 
 | 
   391  | 
  by (blast intro: equalityI)
  | 
| 
 | 
   392  | 
  | 
| 
 | 
   393  | 
lemma Int_assoc: "(A Int B) Int C  =  A Int (B Int C)"
  | 
| 
 | 
   394  | 
  by (blast intro: equalityI)
  | 
| 
 | 
   395  | 
  | 
| 
 | 
   396  | 
lemma Int_Un_distrib: "(A Un B) Int C  =  (A Int C) Un (B Int C)"
  | 
| 
 | 
   397  | 
  by (blast intro: equalityI)
  | 
| 
 | 
   398  | 
  | 
| 
 | 
   399  | 
lemma subset_Int_eq: "(A<=B) <-> (A Int B = A)"
  | 
| 
 | 
   400  | 
  by (blast intro: equalityI elim: equalityE)
  | 
| 
 | 
   401  | 
  | 
| 
 | 
   402  | 
  | 
| 
 | 
   403  | 
subsection {* Binary Union *}
 | 
| 
 | 
   404  | 
  | 
| 
 | 
   405  | 
lemma Un_absorb: "A Un A = A"
  | 
| 
 | 
   406  | 
  by (blast intro: equalityI)
  | 
| 
 | 
   407  | 
  | 
| 
 | 
   408  | 
lemma Un_commute: "A Un B  =  B Un A"
  | 
| 
 | 
   409  | 
  by (blast intro: equalityI)
  | 
| 
 | 
   410  | 
  | 
| 
 | 
   411  | 
lemma Un_assoc: "(A Un B) Un C  =  A Un (B Un C)"
  | 
| 
 | 
   412  | 
  by (blast intro: equalityI)
  | 
| 
 | 
   413  | 
  | 
| 
 | 
   414  | 
lemma Un_Int_distrib: "(A Int B) Un C  =  (A Un C) Int (B Un C)"
  | 
| 
 | 
   415  | 
  by (blast intro: equalityI)
  | 
| 
 | 
   416  | 
  | 
| 
 | 
   417  | 
lemma Un_Int_crazy:
  | 
| 
 | 
   418  | 
    "(A Int B) Un (B Int C) Un (C Int A) = (A Un B) Int (B Un C) Int (C Un A)"
  | 
| 
 | 
   419  | 
  by (blast intro: equalityI)
  | 
| 
 | 
   420  | 
  | 
| 
 | 
   421  | 
lemma subset_Un_eq: "(A<=B) <-> (A Un B = B)"
  | 
| 
 | 
   422  | 
  by (blast intro: equalityI elim: equalityE)
  | 
| 
 | 
   423  | 
  | 
| 
 | 
   424  | 
  | 
| 
 | 
   425  | 
subsection {* Simple properties of @{text "Compl"} -- complement of a set *}
 | 
| 
 | 
   426  | 
  | 
| 
 | 
   427  | 
lemma Compl_disjoint: "A Int Compl(A) = {x. False}"
 | 
| 
 | 
   428  | 
  by (blast intro: equalityI)
  | 
| 
 | 
   429  | 
  | 
| 
 | 
   430  | 
lemma Compl_partition: "A Un Compl(A) = {x. True}"
 | 
| 
 | 
   431  | 
  by (blast intro: equalityI)
  | 
| 
 | 
   432  | 
  | 
| 
 | 
   433  | 
lemma double_complement: "Compl(Compl(A)) = A"
  | 
| 
 | 
   434  | 
  by (blast intro: equalityI)
  | 
| 
 | 
   435  | 
  | 
| 
 | 
   436  | 
lemma Compl_Un: "Compl(A Un B) = Compl(A) Int Compl(B)"
  | 
| 
 | 
   437  | 
  by (blast intro: equalityI)
  | 
| 
 | 
   438  | 
  | 
| 
 | 
   439  | 
lemma Compl_Int: "Compl(A Int B) = Compl(A) Un Compl(B)"
  | 
| 
 | 
   440  | 
  by (blast intro: equalityI)
  | 
| 
 | 
   441  | 
  | 
| 
 | 
   442  | 
lemma Compl_UN: "Compl(UN x:A. B(x)) = (INT x:A. Compl(B(x)))"
  | 
| 
 | 
   443  | 
  by (blast intro: equalityI)
  | 
| 
 | 
   444  | 
  | 
| 
 | 
   445  | 
lemma Compl_INT: "Compl(INT x:A. B(x)) = (UN x:A. Compl(B(x)))"
  | 
| 
 | 
   446  | 
  by (blast intro: equalityI)
  | 
| 
 | 
   447  | 
  | 
| 
 | 
   448  | 
(*Halmos, Naive Set Theory, page 16.*)
  | 
| 
 | 
   449  | 
lemma Un_Int_assoc_eq: "((A Int B) Un C = A Int (B Un C)) <-> (C<=A)"
  | 
| 
 | 
   450  | 
  by (blast intro: equalityI elim: equalityE)
  | 
| 
 | 
   451  | 
  | 
| 
 | 
   452  | 
  | 
| 
 | 
   453  | 
subsection {* Big Union and Intersection *}
 | 
| 
 | 
   454  | 
  | 
| 
 | 
   455  | 
lemma Union_Un_distrib: "Union(A Un B) = Union(A) Un Union(B)"
  | 
| 
 | 
   456  | 
  by (blast intro: equalityI)
  | 
| 
 | 
   457  | 
  | 
| 
 | 
   458  | 
lemma Union_disjoint:
  | 
| 
 | 
   459  | 
    "(Union(C) Int A = {x. False}) <-> (ALL B:C. B Int A = {x. False})"
 | 
| 
 | 
   460  | 
  by (blast intro: equalityI elim: equalityE)
  | 
| 
 | 
   461  | 
  | 
| 
 | 
   462  | 
lemma Inter_Un_distrib: "Inter(A Un B) = Inter(A) Int Inter(B)"
  | 
| 
 | 
   463  | 
  by (blast intro: equalityI)
  | 
| 
 | 
   464  | 
  | 
| 
 | 
   465  | 
  | 
| 
 | 
   466  | 
subsection {* Unions and Intersections of Families *}
 | 
| 
 | 
   467  | 
  | 
| 
 | 
   468  | 
lemma UN_eq: "(UN x:A. B(x)) = Union({Y. EX x:A. Y=B(x)})"
 | 
| 
 | 
   469  | 
  by (blast intro: equalityI)
  | 
| 
 | 
   470  | 
  | 
| 
 | 
   471  | 
(*Look: it has an EXISTENTIAL quantifier*)
  | 
| 
 | 
   472  | 
lemma INT_eq: "(INT x:A. B(x)) = Inter({Y. EX x:A. Y=B(x)})"
 | 
| 
 | 
   473  | 
  by (blast intro: equalityI)
  | 
| 
 | 
   474  | 
  | 
| 
 | 
   475  | 
lemma Int_Union_image: "A Int Union(B) = (UN C:B. A Int C)"
  | 
| 
 | 
   476  | 
  by (blast intro: equalityI)
  | 
| 
 | 
   477  | 
  | 
| 
 | 
   478  | 
lemma Un_Inter_image: "A Un Inter(B) = (INT C:B. A Un C)"
  | 
| 
 | 
   479  | 
  by (blast intro: equalityI)
  | 
| 
 | 
   480  | 
  | 
| 
 | 
   481  | 
  | 
| 
 | 
   482  | 
section {* Monotonicity of various operations *}
 | 
| 
 | 
   483  | 
  | 
| 
 | 
   484  | 
lemma Union_mono: "A<=B ==> Union(A) <= Union(B)"
  | 
| 
 | 
   485  | 
  by blast
  | 
| 
 | 
   486  | 
  | 
| 
 | 
   487  | 
lemma Inter_anti_mono: "[| B<=A |] ==> Inter(A) <= Inter(B)"
  | 
| 
 | 
   488  | 
  by blast
  | 
| 
 | 
   489  | 
  | 
| 
 | 
   490  | 
lemma UN_mono:
  | 
| 
 | 
   491  | 
  "[| A<=B;  !!x. x:A ==> f(x)<=g(x) |] ==>  
  | 
| 
 | 
   492  | 
    (UN x:A. f(x)) <= (UN x:B. g(x))"
  | 
| 
 | 
   493  | 
  by blast
  | 
| 
 | 
   494  | 
  | 
| 
 | 
   495  | 
lemma INT_anti_mono:
  | 
| 
 | 
   496  | 
  "[| B<=A;  !!x. x:A ==> f(x)<=g(x) |] ==>  
  | 
| 
 | 
   497  | 
    (INT x:A. f(x)) <= (INT x:A. g(x))"
  | 
| 
 | 
   498  | 
  by blast
  | 
| 
 | 
   499  | 
  | 
| 
 | 
   500  | 
lemma Un_mono: "[| A<=C;  B<=D |] ==> A Un B <= C Un D"
  | 
| 
 | 
   501  | 
  by blast
  | 
| 
 | 
   502  | 
  | 
| 
 | 
   503  | 
lemma Int_mono: "[| A<=C;  B<=D |] ==> A Int B <= C Int D"
  | 
| 
 | 
   504  | 
  by blast
  | 
| 
 | 
   505  | 
  | 
| 
 | 
   506  | 
lemma Compl_anti_mono: "[| A<=B |] ==> Compl(B) <= Compl(A)"
  | 
| 
 | 
   507  | 
  by blast
  | 
| 
0
 | 
   508  | 
  | 
| 
 | 
   509  | 
end
  |