src/HOL/Library/Set_Idioms.thy
author nipkow
Wed, 03 Oct 2018 09:46:42 +0200
changeset 69107 c2de7a5c8de9
parent 69004 f6a0c8115e9c
child 69313 b021008c5397
permissions -rw-r--r--
shuffle -> shuffles
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
69004
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     1
(*  Title:      HOL/Library/Set_Idioms.thy
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     2
    Author:     Lawrence Paulson (but borrowed from HOL Light)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     3
*)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     4
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     5
section \<open>Set Idioms\<close>
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     6
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     7
theory Set_Idioms
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     8
imports Countable_Set
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     9
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    10
begin
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    11
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    12
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    13
subsection\<open>Idioms for being a suitable union/intersection of something\<close>
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    14
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    15
definition union_of :: "('a set set \<Rightarrow> bool) \<Rightarrow> ('a set \<Rightarrow> bool) \<Rightarrow> 'a set \<Rightarrow> bool"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    16
  (infixr "union'_of" 60)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    17
  where "P union_of Q \<equiv> \<lambda>S. \<exists>\<U>. P \<U> \<and> \<U> \<subseteq> Collect Q \<and> \<Union>\<U> = S"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    18
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    19
definition intersection_of :: "('a set set \<Rightarrow> bool) \<Rightarrow> ('a set \<Rightarrow> bool) \<Rightarrow> 'a set \<Rightarrow> bool"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    20
  (infixr "intersection'_of" 60)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    21
  where "P intersection_of Q \<equiv> \<lambda>S. \<exists>\<U>. P \<U> \<and> \<U> \<subseteq> Collect Q \<and> \<Inter>\<U> = S"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    22
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    23
definition arbitrary:: "'a set set \<Rightarrow> bool" where "arbitrary \<U> \<equiv> True"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    24
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    25
lemma union_of_inc: "\<lbrakk>P {S}; Q S\<rbrakk> \<Longrightarrow> (P union_of Q) S"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    26
  by (auto simp: union_of_def)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    27
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    28
lemma intersection_of_inc:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    29
   "\<lbrakk>P {S}; Q S\<rbrakk> \<Longrightarrow> (P intersection_of Q) S"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    30
  by (auto simp: intersection_of_def)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    31
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    32
lemma union_of_mono:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    33
   "\<lbrakk>(P union_of Q) S; \<And>x. Q x \<Longrightarrow> Q' x\<rbrakk> \<Longrightarrow> (P union_of Q') S"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    34
  by (auto simp: union_of_def)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    35
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    36
lemma intersection_of_mono:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    37
   "\<lbrakk>(P intersection_of Q) S; \<And>x. Q x \<Longrightarrow> Q' x\<rbrakk> \<Longrightarrow> (P intersection_of Q') S"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    38
  by (auto simp: intersection_of_def)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    39
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    40
lemma all_union_of:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    41
     "(\<forall>S. (P union_of Q) S \<longrightarrow> R S) \<longleftrightarrow> (\<forall>T. P T \<and> T \<subseteq> Collect Q \<longrightarrow> R(\<Union>T))"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    42
  by (auto simp: union_of_def)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    43
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    44
lemma all_intersection_of:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    45
     "(\<forall>S. (P intersection_of Q) S \<longrightarrow> R S) \<longleftrightarrow> (\<forall>T. P T \<and> T \<subseteq> Collect Q \<longrightarrow> R(\<Inter>T))"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    46
  by (auto simp: intersection_of_def)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    47
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    48
lemma union_of_empty:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    49
     "P {} \<Longrightarrow> (P union_of Q) {}"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    50
  by (auto simp: union_of_def)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    51
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    52
lemma intersection_of_empty:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    53
     "P {} \<Longrightarrow> (P intersection_of Q) UNIV"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    54
  by (auto simp: intersection_of_def)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    55
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    56
text\<open> The arbitrary and finite cases\<close>
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    57
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    58
lemma arbitrary_union_of_alt:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    59
   "(arbitrary union_of Q) S \<longleftrightarrow> (\<forall>x \<in> S. \<exists>U. Q U \<and> x \<in> U \<and> U \<subseteq> S)"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    60
 (is "?lhs = ?rhs")
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    61
proof
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    62
  assume ?lhs
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    63
  then show ?rhs
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    64
    by (force simp: union_of_def arbitrary_def)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    65
next
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    66
  assume ?rhs
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    67
  then have "{U. Q U \<and> U \<subseteq> S} \<subseteq> Collect Q" "\<Union>{U. Q U \<and> U \<subseteq> S} = S"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    68
    by auto
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    69
  then show ?lhs
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    70
    unfolding union_of_def arbitrary_def by blast
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    71
qed
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    72
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    73
lemma arbitrary_union_of_empty [simp]: "(arbitrary union_of P) {}"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    74
  by (force simp: union_of_def arbitrary_def)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    75
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    76
lemma arbitrary_intersection_of_empty [simp]:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    77
  "(arbitrary intersection_of P) UNIV"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    78
  by (force simp: intersection_of_def arbitrary_def)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    79
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    80
lemma arbitrary_union_of_inc:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    81
  "P S \<Longrightarrow> (arbitrary union_of P) S"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    82
  by (force simp: union_of_inc arbitrary_def)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    83
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    84
lemma arbitrary_intersection_of_inc:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    85
  "P S \<Longrightarrow> (arbitrary intersection_of P) S"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    86
  by (force simp: intersection_of_inc arbitrary_def)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    87
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    88
lemma arbitrary_union_of_complement:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    89
   "(arbitrary union_of P) S \<longleftrightarrow> (arbitrary intersection_of (\<lambda>S. P(- S))) (- S)"  (is "?lhs = ?rhs")
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    90
proof
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    91
  assume ?lhs
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    92
  then obtain \<U> where "\<U> \<subseteq> Collect P" "S = \<Union>\<U>"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    93
    by (auto simp: union_of_def arbitrary_def)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    94
  then show ?rhs
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    95
    unfolding intersection_of_def arbitrary_def
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    96
    by (rule_tac x="uminus ` \<U>" in exI) auto
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    97
next
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    98
  assume ?rhs
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    99
  then obtain \<U> where "\<U> \<subseteq> {S. P (- S)}" "\<Inter>\<U> = - S"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   100
    by (auto simp: union_of_def intersection_of_def arbitrary_def)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   101
  then show ?lhs
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   102
    unfolding union_of_def arbitrary_def
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   103
    by (rule_tac x="uminus ` \<U>" in exI) auto
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   104
qed
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   105
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   106
lemma arbitrary_intersection_of_complement:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   107
   "(arbitrary intersection_of P) S \<longleftrightarrow> (arbitrary union_of (\<lambda>S. P(- S))) (- S)"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   108
  by (simp add: arbitrary_union_of_complement)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   109
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   110
lemma arbitrary_union_of_idempot [simp]:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   111
  "arbitrary union_of arbitrary union_of P = arbitrary union_of P"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   112
proof -
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   113
  have 1: "\<exists>\<U>'\<subseteq>Collect P. \<Union>\<U>' = \<Union>\<U>" if "\<U> \<subseteq> {S. \<exists>\<V>\<subseteq>Collect P. \<Union>\<V> = S}" for \<U>
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   114
  proof -
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   115
    let ?\<W> = "{V. \<exists>\<V>. \<V>\<subseteq>Collect P \<and> V \<in> \<V> \<and> (\<exists>S \<in> \<U>. \<Union>\<V> = S)}"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   116
    have *: "\<And>x U. \<lbrakk>x \<in> U; U \<in> \<U>\<rbrakk> \<Longrightarrow> x \<in> \<Union>?\<W>"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   117
      using that
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   118
      apply simp
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   119
      apply (drule subsetD, assumption, auto)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   120
      done
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   121
    show ?thesis
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   122
    apply (rule_tac x="{V. \<exists>\<V>. \<V>\<subseteq>Collect P \<and> V \<in> \<V> \<and> (\<exists>S \<in> \<U>. \<Union>\<V> = S)}" in exI)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   123
      using that by (blast intro: *)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   124
  qed
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   125
  have 2: "\<exists>\<U>'\<subseteq>{S. \<exists>\<U>\<subseteq>Collect P. \<Union>\<U> = S}. \<Union>\<U>' = \<Union>\<U>" if "\<U> \<subseteq> Collect P" for \<U>
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   126
    by (metis (mono_tags, lifting) union_of_def arbitrary_union_of_inc that)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   127
  show ?thesis
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   128
    unfolding union_of_def arbitrary_def by (force simp: 1 2)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   129
qed
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   130
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   131
lemma arbitrary_intersection_of_idempot:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   132
   "arbitrary intersection_of arbitrary intersection_of P = arbitrary intersection_of P" (is "?lhs = ?rhs")
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   133
proof -
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   134
  have "- ?lhs = - ?rhs"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   135
    unfolding arbitrary_intersection_of_complement by simp
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   136
  then show ?thesis
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   137
    by simp
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   138
qed
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   139
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   140
lemma arbitrary_union_of_Union:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   141
   "(\<And>S. S \<in> \<U> \<Longrightarrow> (arbitrary union_of P) S) \<Longrightarrow> (arbitrary union_of P) (\<Union>\<U>)"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   142
  by (metis union_of_def arbitrary_def arbitrary_union_of_idempot mem_Collect_eq subsetI)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   143
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   144
lemma arbitrary_union_of_Un:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   145
   "\<lbrakk>(arbitrary union_of P) S; (arbitrary union_of P) T\<rbrakk>
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   146
           \<Longrightarrow> (arbitrary union_of P) (S \<union> T)"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   147
  using arbitrary_union_of_Union [of "{S,T}"] by auto
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   148
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   149
lemma arbitrary_intersection_of_Inter:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   150
   "(\<And>S. S \<in> \<U> \<Longrightarrow> (arbitrary intersection_of P) S) \<Longrightarrow> (arbitrary intersection_of P) (\<Inter>\<U>)"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   151
  by (metis intersection_of_def arbitrary_def arbitrary_intersection_of_idempot mem_Collect_eq subsetI)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   152
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   153
lemma arbitrary_intersection_of_Int:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   154
   "\<lbrakk>(arbitrary intersection_of P) S; (arbitrary intersection_of P) T\<rbrakk>
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   155
           \<Longrightarrow> (arbitrary intersection_of P) (S \<inter> T)"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   156
  using arbitrary_intersection_of_Inter [of "{S,T}"] by auto
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   157
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   158
lemma arbitrary_union_of_Int_eq:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   159
  "(\<forall>S T. (arbitrary union_of P) S \<and> (arbitrary union_of P) T
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   160
               \<longrightarrow> (arbitrary union_of P) (S \<inter> T))
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   161
   \<longleftrightarrow> (\<forall>S T. P S \<and> P T \<longrightarrow> (arbitrary union_of P) (S \<inter> T))"  (is "?lhs = ?rhs")
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   162
proof
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   163
  assume ?lhs
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   164
  then show ?rhs
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   165
    by (simp add: arbitrary_union_of_inc)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   166
next
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   167
  assume R: ?rhs
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   168
  show ?lhs
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   169
  proof clarify
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   170
    fix S :: "'a set" and T :: "'a set"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   171
    assume "(arbitrary union_of P) S" and "(arbitrary union_of P) T"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   172
    then obtain \<U> \<V> where *: "\<U> \<subseteq> Collect P" "\<Union>\<U> = S" "\<V> \<subseteq> Collect P" "\<Union>\<V> = T"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   173
      by (auto simp: union_of_def)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   174
    then have "(arbitrary union_of P) (\<Union>C\<in>\<U>. \<Union>D\<in>\<V>. C \<inter> D)"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   175
     using R by (blast intro: arbitrary_union_of_Union)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   176
   then show "(arbitrary union_of P) (S \<inter> T)"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   177
     by (simp add: Int_UN_distrib2 *)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   178
 qed
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   179
qed
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   180
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   181
lemma arbitrary_intersection_of_Un_eq:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   182
   "(\<forall>S T. (arbitrary intersection_of P) S \<and> (arbitrary intersection_of P) T
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   183
               \<longrightarrow> (arbitrary intersection_of P) (S \<union> T)) \<longleftrightarrow>
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   184
        (\<forall>S T. P S \<and> P T \<longrightarrow> (arbitrary intersection_of P) (S \<union> T))"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   185
  apply (simp add: arbitrary_intersection_of_complement)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   186
  using arbitrary_union_of_Int_eq [of "\<lambda>S. P (- S)"]
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   187
  by (metis (no_types, lifting) arbitrary_def double_compl union_of_inc)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   188
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   189
lemma finite_union_of_empty [simp]: "(finite union_of P) {}"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   190
  by (simp add: union_of_empty)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   191
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   192
lemma finite_intersection_of_empty [simp]: "(finite intersection_of P) UNIV"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   193
  by (simp add: intersection_of_empty)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   194
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   195
lemma finite_union_of_inc:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   196
   "P S \<Longrightarrow> (finite union_of P) S"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   197
  by (simp add: union_of_inc)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   198
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   199
lemma finite_intersection_of_inc:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   200
   "P S \<Longrightarrow> (finite intersection_of P) S"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   201
  by (simp add: intersection_of_inc)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   202
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   203
lemma finite_union_of_complement:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   204
  "(finite union_of P) S \<longleftrightarrow> (finite intersection_of (\<lambda>S. P(- S))) (- S)"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   205
  unfolding union_of_def intersection_of_def
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   206
  apply safe
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   207
   apply (rule_tac x="uminus ` \<U>" in exI, fastforce)+
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   208
  done
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   209
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   210
lemma finite_intersection_of_complement:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   211
   "(finite intersection_of P) S \<longleftrightarrow> (finite union_of (\<lambda>S. P(- S))) (- S)"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   212
  by (simp add: finite_union_of_complement)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   213
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   214
lemma finite_union_of_idempot [simp]:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   215
  "finite union_of finite union_of P = finite union_of P"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   216
proof -
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   217
  have "(finite union_of P) S" if S: "(finite union_of finite union_of P) S" for S
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   218
  proof -
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   219
    obtain \<U> where "finite \<U>" "S = \<Union>\<U>" and \<U>: "\<forall>U\<in>\<U>. \<exists>\<U>. finite \<U> \<and> (\<U> \<subseteq> Collect P) \<and> \<Union>\<U> = U"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   220
      using S unfolding union_of_def by (auto simp: subset_eq)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   221
    then obtain f where "\<forall>U\<in>\<U>. finite (f U) \<and> (f U \<subseteq> Collect P) \<and> \<Union>(f U) = U"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   222
      by metis
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   223
    then show ?thesis
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   224
      unfolding union_of_def \<open>S = \<Union>\<U>\<close>
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   225
      by (rule_tac x = "snd ` Sigma \<U> f" in exI) (fastforce simp: \<open>finite \<U>\<close>)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   226
  qed
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   227
  moreover
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   228
  have "(finite union_of finite union_of P) S" if "(finite union_of P) S" for S
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   229
    by (simp add: finite_union_of_inc that)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   230
  ultimately show ?thesis
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   231
    by force
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   232
qed
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   233
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   234
lemma finite_intersection_of_idempot [simp]:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   235
   "finite intersection_of finite intersection_of P = finite intersection_of P"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   236
  by (force simp: finite_intersection_of_complement)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   237
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   238
lemma finite_union_of_Union:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   239
   "\<lbrakk>finite \<U>; \<And>S. S \<in> \<U> \<Longrightarrow> (finite union_of P) S\<rbrakk> \<Longrightarrow> (finite union_of P) (\<Union>\<U>)"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   240
  using finite_union_of_idempot [of P]
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   241
  by (metis mem_Collect_eq subsetI union_of_def)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   242
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   243
lemma finite_union_of_Un:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   244
   "\<lbrakk>(finite union_of P) S; (finite union_of P) T\<rbrakk> \<Longrightarrow> (finite union_of P) (S \<union> T)"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   245
  by (auto simp: union_of_def)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   246
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   247
lemma finite_intersection_of_Inter:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   248
   "\<lbrakk>finite \<U>; \<And>S. S \<in> \<U> \<Longrightarrow> (finite intersection_of P) S\<rbrakk> \<Longrightarrow> (finite intersection_of P) (\<Inter>\<U>)"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   249
  using finite_intersection_of_idempot [of P]
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   250
  by (metis intersection_of_def mem_Collect_eq subsetI)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   251
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   252
lemma finite_intersection_of_Int:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   253
   "\<lbrakk>(finite intersection_of P) S; (finite intersection_of P) T\<rbrakk>
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   254
           \<Longrightarrow> (finite intersection_of P) (S \<inter> T)"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   255
  by (auto simp: intersection_of_def)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   256
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   257
lemma finite_union_of_Int_eq:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   258
   "(\<forall>S T. (finite union_of P) S \<and> (finite union_of P) T \<longrightarrow> (finite union_of P) (S \<inter> T))
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   259
    \<longleftrightarrow> (\<forall>S T. P S \<and> P T \<longrightarrow> (finite union_of P) (S \<inter> T))"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   260
(is "?lhs = ?rhs")
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   261
proof
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   262
  assume ?lhs
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   263
  then show ?rhs
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   264
    by (simp add: finite_union_of_inc)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   265
next
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   266
  assume R: ?rhs
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   267
  show ?lhs
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   268
  proof clarify
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   269
    fix S :: "'a set" and T :: "'a set"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   270
    assume "(finite union_of P) S" and "(finite union_of P) T"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   271
    then obtain \<U> \<V> where *: "\<U> \<subseteq> Collect P" "\<Union>\<U> = S" "finite \<U>" "\<V> \<subseteq> Collect P" "\<Union>\<V> = T" "finite \<V>"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   272
      by (auto simp: union_of_def)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   273
    then have "(finite union_of P) (\<Union>C\<in>\<U>. \<Union>D\<in>\<V>. C \<inter> D)"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   274
      using R
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   275
      by (blast intro: finite_union_of_Union)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   276
   then show "(finite union_of P) (S \<inter> T)"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   277
     by (simp add: Int_UN_distrib2 *)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   278
 qed
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   279
qed
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   280
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   281
lemma finite_intersection_of_Un_eq:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   282
   "(\<forall>S T. (finite intersection_of P) S \<and>
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   283
               (finite intersection_of P) T
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   284
               \<longrightarrow> (finite intersection_of P) (S \<union> T)) \<longleftrightarrow>
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   285
        (\<forall>S T. P S \<and> P T \<longrightarrow> (finite intersection_of P) (S \<union> T))"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   286
  apply (simp add: finite_intersection_of_complement)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   287
  using finite_union_of_Int_eq [of "\<lambda>S. P (- S)"]
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   288
  by (metis (no_types, lifting) double_compl)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   289
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   290
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   291
abbreviation finite' :: "'a set \<Rightarrow> bool"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   292
  where "finite' A \<equiv> finite A \<and> A \<noteq> {}"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   293
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   294
lemma finite'_intersection_of_Int:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   295
   "\<lbrakk>(finite' intersection_of P) S; (finite' intersection_of P) T\<rbrakk>
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   296
           \<Longrightarrow> (finite' intersection_of P) (S \<inter> T)"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   297
  by (auto simp: intersection_of_def)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   298
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   299
lemma finite'_intersection_of_inc:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   300
   "P S \<Longrightarrow> (finite' intersection_of P) S"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   301
  by (simp add: intersection_of_inc)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   302
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   303
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   304
subsection \<open>The ``Relative to'' operator\<close>
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   305
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   306
text\<open>A somewhat cheap but handy way of getting localized forms of various topological concepts
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   307
(open, closed, borel, fsigma, gdelta etc.)\<close>
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   308
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   309
definition relative_to :: "['a set \<Rightarrow> bool, 'a set, 'a set] \<Rightarrow> bool" (infixl "relative'_to" 55)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   310
  where "P relative_to S \<equiv> \<lambda>T. \<exists>U. P U \<and> S \<inter> U = T"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   311
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   312
lemma relative_to_UNIV [simp]: "(P relative_to UNIV) S \<longleftrightarrow> P S"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   313
  by (simp add: relative_to_def)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   314
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   315
lemma relative_to_imp_subset:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   316
   "(P relative_to S) T \<Longrightarrow> T \<subseteq> S"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   317
  by (auto simp: relative_to_def)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   318
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   319
lemma all_relative_to: "(\<forall>S. (P relative_to U) S \<longrightarrow> Q S) \<longleftrightarrow> (\<forall>S. P S \<longrightarrow> Q(U \<inter> S))"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   320
  by (auto simp: relative_to_def)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   321
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   322
lemma relative_to_inc:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   323
   "P S \<Longrightarrow> (P relative_to U) (U \<inter> S)"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   324
  by (auto simp: relative_to_def)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   325
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   326
lemma relative_to_relative_to [simp]:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   327
   "P relative_to S relative_to T = P relative_to (S \<inter> T)"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   328
  unfolding relative_to_def
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   329
  by auto
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   330
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   331
lemma relative_to_compl:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   332
   "S \<subseteq> U \<Longrightarrow> ((P relative_to U) (U - S) \<longleftrightarrow> ((\<lambda>c. P(- c)) relative_to U) S)"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   333
  unfolding relative_to_def
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   334
  by (metis Diff_Diff_Int Diff_eq double_compl inf.absorb_iff2)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   335
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   336
lemma relative_to_subset:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   337
   "S \<subseteq> T \<and> P S \<Longrightarrow> (P relative_to T) S"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   338
  unfolding relative_to_def by auto
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   339
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   340
lemma relative_to_subset_trans:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   341
   "(P relative_to U) S \<and> S \<subseteq> T \<and> T \<subseteq> U \<Longrightarrow> (P relative_to T) S"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   342
  unfolding relative_to_def by auto
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   343
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   344
lemma relative_to_mono:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   345
   "\<lbrakk>(P relative_to U) S; \<And>S. P S \<Longrightarrow> Q S\<rbrakk> \<Longrightarrow> (Q relative_to U) S"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   346
  unfolding relative_to_def by auto
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   347
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   348
lemma relative_to_subset_inc: "\<lbrakk>S \<subseteq> U; P S\<rbrakk> \<Longrightarrow> (P relative_to U) S"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   349
  unfolding relative_to_def by auto
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   350
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   351
lemma relative_to_Int:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   352
   "\<lbrakk>(P relative_to S) C; (P relative_to S) D; \<And>X Y. \<lbrakk>P X; P Y\<rbrakk> \<Longrightarrow> P(X \<inter> Y)\<rbrakk>
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   353
         \<Longrightarrow>  (P relative_to S) (C \<inter> D)"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   354
  unfolding relative_to_def by auto
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   355
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   356
lemma relative_to_Un:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   357
   "\<lbrakk>(P relative_to S) C; (P relative_to S) D; \<And>X Y. \<lbrakk>P X; P Y\<rbrakk> \<Longrightarrow> P(X \<union> Y)\<rbrakk>
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   358
         \<Longrightarrow>  (P relative_to S) (C \<union> D)"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   359
  unfolding relative_to_def by auto
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   360
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   361
lemma arbitrary_union_of_relative_to:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   362
  "((arbitrary union_of P) relative_to U) = (arbitrary union_of (P relative_to U))" (is "?lhs = ?rhs")
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   363
proof -
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   364
  have "?rhs S" if L: "?lhs S" for S
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   365
  proof -
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   366
    obtain \<U> where "S = U \<inter> \<Union>\<U>" "\<U> \<subseteq> Collect P"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   367
      using L unfolding relative_to_def union_of_def by auto
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   368
    then show ?thesis
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   369
      unfolding relative_to_def union_of_def arbitrary_def
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   370
      by (rule_tac x="(\<lambda>X. U \<inter> X) ` \<U>" in exI) auto
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   371
  qed
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   372
  moreover have "?lhs S" if R: "?rhs S" for S
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   373
  proof -
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   374
    obtain \<U> where "S = \<Union>\<U>" "\<forall>T\<in>\<U>. \<exists>V. P V \<and> U \<inter> V = T"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   375
      using R unfolding relative_to_def union_of_def by auto
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   376
    then obtain f where f: "\<And>T. T \<in> \<U> \<Longrightarrow> P (f T)" "\<And>T. T \<in> \<U> \<Longrightarrow> U \<inter> (f T) = T"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   377
      by metis
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   378
    then have "\<exists>\<U>'\<subseteq>Collect P. \<Union>\<U>' = UNION \<U> f"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   379
      by (metis image_subset_iff mem_Collect_eq)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   380
    moreover have eq: "U \<inter> UNION \<U> f = \<Union>\<U>"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   381
      using f by auto
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   382
    ultimately show ?thesis
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   383
      unfolding relative_to_def union_of_def arbitrary_def \<open>S = \<Union>\<U>\<close>
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   384
      by metis
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   385
  qed
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   386
  ultimately show ?thesis
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   387
    by blast
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   388
qed
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   389
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   390
lemma finite_union_of_relative_to:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   391
  "((finite union_of P) relative_to U) = (finite union_of (P relative_to U))" (is "?lhs = ?rhs")
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   392
proof -
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   393
  have "?rhs S" if L: "?lhs S" for S
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   394
  proof -
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   395
    obtain \<U> where "S = U \<inter> \<Union>\<U>" "\<U> \<subseteq> Collect P" "finite \<U>"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   396
      using L unfolding relative_to_def union_of_def by auto
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   397
    then show ?thesis
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   398
      unfolding relative_to_def union_of_def
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   399
      by (rule_tac x="(\<lambda>X. U \<inter> X) ` \<U>" in exI) auto
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   400
  qed
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   401
  moreover have "?lhs S" if R: "?rhs S" for S
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   402
  proof -
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   403
    obtain \<U> where "S = \<Union>\<U>" "\<forall>T\<in>\<U>. \<exists>V. P V \<and> U \<inter> V = T" "finite \<U>"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   404
      using R unfolding relative_to_def union_of_def by auto
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   405
    then obtain f where f: "\<And>T. T \<in> \<U> \<Longrightarrow> P (f T)" "\<And>T. T \<in> \<U> \<Longrightarrow> U \<inter> (f T) = T"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   406
      by metis
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   407
    then have "\<exists>\<U>'\<subseteq>Collect P. \<Union>\<U>' = UNION \<U> f"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   408
      by (metis image_subset_iff mem_Collect_eq)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   409
    moreover have eq: "U \<inter> UNION \<U> f = \<Union>\<U>"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   410
      using f by auto
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   411
    ultimately show ?thesis
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   412
      using \<open>finite \<U>\<close> f
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   413
      unfolding relative_to_def union_of_def \<open>S = \<Union>\<U>\<close>
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   414
      by (rule_tac x="UNION \<U> f" in exI) (metis finite_imageI image_subsetI mem_Collect_eq)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   415
  qed
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   416
  ultimately show ?thesis
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   417
    by blast
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   418
qed
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   419
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   420
lemma countable_union_of_relative_to:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   421
  "((countable union_of P) relative_to U) = (countable union_of (P relative_to U))" (is "?lhs = ?rhs")
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   422
proof -
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   423
  have "?rhs S" if L: "?lhs S" for S
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   424
  proof -
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   425
    obtain \<U> where "S = U \<inter> \<Union>\<U>" "\<U> \<subseteq> Collect P" "countable \<U>"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   426
      using L unfolding relative_to_def union_of_def by auto
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   427
    then show ?thesis
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   428
      unfolding relative_to_def union_of_def
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   429
      by (rule_tac x="(\<lambda>X. U \<inter> X) ` \<U>" in exI) auto
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   430
  qed
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   431
  moreover have "?lhs S" if R: "?rhs S" for S
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   432
  proof -
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   433
    obtain \<U> where "S = \<Union>\<U>" "\<forall>T\<in>\<U>. \<exists>V. P V \<and> U \<inter> V = T" "countable \<U>"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   434
      using R unfolding relative_to_def union_of_def by auto
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   435
    then obtain f where f: "\<And>T. T \<in> \<U> \<Longrightarrow> P (f T)" "\<And>T. T \<in> \<U> \<Longrightarrow> U \<inter> (f T) = T"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   436
      by metis
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   437
    then have "\<exists>\<U>'\<subseteq>Collect P. \<Union>\<U>' = UNION \<U> f"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   438
      by (metis image_subset_iff mem_Collect_eq)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   439
    moreover have eq: "U \<inter> UNION \<U> f = \<Union>\<U>"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   440
      using f by auto
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   441
    ultimately show ?thesis
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   442
      using \<open>countable \<U>\<close> f
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   443
      unfolding relative_to_def union_of_def \<open>S = \<Union>\<U>\<close>
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   444
      by (rule_tac x="UNION \<U> f" in exI) (metis countable_image image_subsetI mem_Collect_eq)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   445
  qed
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   446
  ultimately show ?thesis
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   447
    by blast
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   448
qed
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   449
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   450
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   451
lemma arbitrary_intersection_of_relative_to:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   452
  "((arbitrary intersection_of P) relative_to U) = ((arbitrary intersection_of (P relative_to U)) relative_to U)" (is "?lhs = ?rhs")
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   453
proof -
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   454
  have "?rhs S" if L: "?lhs S" for S
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   455
  proof -
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   456
    obtain \<U> where \<U>: "S = U \<inter> \<Inter>\<U>" "\<U> \<subseteq> Collect P"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   457
      using L unfolding relative_to_def intersection_of_def by auto
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   458
    show ?thesis
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   459
      unfolding relative_to_def intersection_of_def arbitrary_def
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   460
    proof (intro exI conjI)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   461
      show "U \<inter> (\<Inter>X\<in>\<U>. U \<inter> X) = S" "(\<inter>) U ` \<U> \<subseteq> {T. \<exists>Ua. P Ua \<and> U \<inter> Ua = T}"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   462
        using \<U> by blast+
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   463
    qed auto
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   464
  qed
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   465
  moreover have "?lhs S" if R: "?rhs S" for S
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   466
  proof -
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   467
    obtain \<U> where "S = U \<inter> \<Inter>\<U>" "\<forall>T\<in>\<U>. \<exists>V. P V \<and> U \<inter> V = T"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   468
      using R unfolding relative_to_def intersection_of_def  by auto
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   469
    then obtain f where f: "\<And>T. T \<in> \<U> \<Longrightarrow> P (f T)" "\<And>T. T \<in> \<U> \<Longrightarrow> U \<inter> (f T) = T"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   470
      by metis
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   471
    then have "\<exists>\<U>'\<subseteq>Collect P. \<Inter>\<U>' = INTER \<U> f"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   472
      by (metis image_subset_iff mem_Collect_eq)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   473
    moreover have eq: "U \<inter> INTER \<U> f = U \<inter> \<Inter>\<U>"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   474
      using f by auto
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   475
    ultimately show ?thesis
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   476
      unfolding relative_to_def intersection_of_def arbitrary_def \<open>S = U \<inter> \<Inter>\<U>\<close>
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   477
      by metis
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   478
  qed
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   479
  ultimately show ?thesis
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   480
    by blast
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   481
qed
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   482
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   483
lemma finite_intersection_of_relative_to:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   484
  "((finite intersection_of P) relative_to U) = ((finite intersection_of (P relative_to U)) relative_to U)" (is "?lhs = ?rhs")
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   485
proof -
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   486
  have "?rhs S" if L: "?lhs S" for S
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   487
  proof -
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   488
    obtain \<U> where \<U>: "S = U \<inter> \<Inter>\<U>" "\<U> \<subseteq> Collect P" "finite \<U>"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   489
      using L unfolding relative_to_def intersection_of_def by auto
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   490
    show ?thesis
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   491
      unfolding relative_to_def intersection_of_def
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   492
    proof (intro exI conjI)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   493
      show "U \<inter> (\<Inter>X\<in>\<U>. U \<inter> X) = S" "(\<inter>) U ` \<U> \<subseteq> {T. \<exists>Ua. P Ua \<and> U \<inter> Ua = T}"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   494
        using \<U> by blast+
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   495
      show "finite ((\<inter>) U ` \<U>)"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   496
        by (simp add: \<open>finite \<U>\<close>)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   497
    qed auto
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   498
  qed
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   499
  moreover have "?lhs S" if R: "?rhs S" for S
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   500
  proof -
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   501
    obtain \<U> where "S = U \<inter> \<Inter>\<U>" "\<forall>T\<in>\<U>. \<exists>V. P V \<and> U \<inter> V = T" "finite \<U>"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   502
      using R unfolding relative_to_def intersection_of_def  by auto
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   503
    then obtain f where f: "\<And>T. T \<in> \<U> \<Longrightarrow> P (f T)" "\<And>T. T \<in> \<U> \<Longrightarrow> U \<inter> (f T) = T"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   504
      by metis
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   505
    then have "\<exists>\<U>'\<subseteq>Collect P. \<Inter>\<U>' = INTER \<U> f"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   506
      by (metis image_subset_iff mem_Collect_eq)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   507
    moreover have eq: "U \<inter> INTER \<U> f = U \<inter> \<Inter>\<U>"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   508
      using f by auto
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   509
    ultimately show ?thesis
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   510
      unfolding relative_to_def intersection_of_def \<open>S = U \<inter> \<Inter>\<U>\<close>
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   511
      by (metis (no_types, hide_lams) \<open>finite \<U>\<close> f(1) finite_imageI image_subset_iff mem_Collect_eq)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   512
  qed
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   513
  ultimately show ?thesis
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   514
    by blast
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   515
qed
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   516
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   517
lemma countable_intersection_of_relative_to:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   518
  "((countable intersection_of P) relative_to U) = ((countable intersection_of (P relative_to U)) relative_to U)" (is "?lhs = ?rhs")
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   519
proof -
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   520
  have "?rhs S" if L: "?lhs S" for S
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   521
  proof -
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   522
    obtain \<U> where \<U>: "S = U \<inter> \<Inter>\<U>" "\<U> \<subseteq> Collect P" "countable \<U>"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   523
      using L unfolding relative_to_def intersection_of_def by auto
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   524
    show ?thesis
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   525
      unfolding relative_to_def intersection_of_def
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   526
    proof (intro exI conjI)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   527
      show "U \<inter> (\<Inter>X\<in>\<U>. U \<inter> X) = S" "(\<inter>) U ` \<U> \<subseteq> {T. \<exists>Ua. P Ua \<and> U \<inter> Ua = T}"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   528
        using \<U> by blast+
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   529
      show "countable ((\<inter>) U ` \<U>)"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   530
        by (simp add: \<open>countable \<U>\<close>)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   531
    qed auto
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   532
  qed
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   533
  moreover have "?lhs S" if R: "?rhs S" for S
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   534
  proof -
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   535
    obtain \<U> where "S = U \<inter> \<Inter>\<U>" "\<forall>T\<in>\<U>. \<exists>V. P V \<and> U \<inter> V = T" "countable \<U>"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   536
      using R unfolding relative_to_def intersection_of_def  by auto
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   537
    then obtain f where f: "\<And>T. T \<in> \<U> \<Longrightarrow> P (f T)" "\<And>T. T \<in> \<U> \<Longrightarrow> U \<inter> (f T) = T"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   538
      by metis
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   539
    then have "\<exists>\<U>'\<subseteq>Collect P. \<Inter>\<U>' = INTER \<U> f"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   540
      by (metis image_subset_iff mem_Collect_eq)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   541
    moreover have eq: "U \<inter> INTER \<U> f = U \<inter> \<Inter>\<U>"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   542
      using f by auto
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   543
    ultimately show ?thesis
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   544
      unfolding relative_to_def intersection_of_def \<open>S = U \<inter> \<Inter>\<U>\<close>
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   545
      by (metis (no_types, hide_lams) \<open>countable \<U>\<close> f(1) countable_image image_subset_iff mem_Collect_eq)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   546
  qed
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   547
  ultimately show ?thesis
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   548
    by blast
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   549
qed
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   550
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   551
end