src/HOL/Library/Set_Idioms.thy
author wenzelm
Tue, 05 Nov 2019 14:16:16 +0100
changeset 71046 b8aeeedf7e68
parent 69918 eddcc7c726f3
child 77935 7f240b0dabd9
permissions -rw-r--r--
support for Linux user management;
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)
69918
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69325
diff changeset
    47
             
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69325
diff changeset
    48
lemma intersection_ofE:
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69325
diff changeset
    49
     "\<lbrakk>(P intersection_of Q) S; \<And>T. \<lbrakk>P T; T \<subseteq> Collect Q\<rbrakk> \<Longrightarrow> R(\<Inter>T)\<rbrakk> \<Longrightarrow> R S"
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69325
diff changeset
    50
  by (auto simp: intersection_of_def)
69004
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 union_of_empty:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    53
     "P {} \<Longrightarrow> (P union_of Q) {}"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    54
  by (auto simp: union_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
lemma intersection_of_empty:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    57
     "P {} \<Longrightarrow> (P intersection_of Q) UNIV"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    58
  by (auto simp: intersection_of_def)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    59
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    60
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
    61
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    62
lemma arbitrary_union_of_alt:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    63
   "(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
    64
 (is "?lhs = ?rhs")
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    65
proof
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    66
  assume ?lhs
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    67
  then show ?rhs
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    68
    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
    69
next
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    70
  assume ?rhs
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    71
  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
    72
    by auto
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    73
  then show ?lhs
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    74
    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
    75
qed
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    76
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    77
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
    78
  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
    79
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    80
lemma arbitrary_intersection_of_empty [simp]:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    81
  "(arbitrary intersection_of P) UNIV"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    82
  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
    83
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    84
lemma arbitrary_union_of_inc:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    85
  "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
    86
  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
    87
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    88
lemma arbitrary_intersection_of_inc:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    89
  "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
    90
  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
    91
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    92
lemma arbitrary_union_of_complement:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    93
   "(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
    94
proof
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    95
  assume ?lhs
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    96
  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
    97
    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
    98
  then show ?rhs
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    99
    unfolding intersection_of_def arbitrary_def
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   100
    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
   101
next
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   102
  assume ?rhs
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   103
  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
   104
    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
   105
  then show ?lhs
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   106
    unfolding union_of_def arbitrary_def
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   107
    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
   108
qed
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_intersection_of_complement:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   111
   "(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
   112
  by (simp add: arbitrary_union_of_complement)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   113
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   114
lemma arbitrary_union_of_idempot [simp]:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   115
  "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
   116
proof -
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   117
  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
   118
  proof -
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   119
    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
   120
    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
   121
      using that
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   122
      apply simp
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   123
      apply (drule subsetD, assumption, auto)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   124
      done
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   125
    show ?thesis
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   126
    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
   127
      using that by (blast intro: *)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   128
  qed
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   129
  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
   130
    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
   131
  show ?thesis
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   132
    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
   133
qed
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   134
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   135
lemma arbitrary_intersection_of_idempot:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   136
   "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
   137
proof -
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   138
  have "- ?lhs = - ?rhs"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   139
    unfolding arbitrary_intersection_of_complement by simp
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   140
  then show ?thesis
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   141
    by simp
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   142
qed
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_Union:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   145
   "(\<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
   146
  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
   147
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   148
lemma arbitrary_union_of_Un:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   149
   "\<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
   150
           \<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
   151
  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
   152
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   153
lemma arbitrary_intersection_of_Inter:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   154
   "(\<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
   155
  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
   156
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   157
lemma arbitrary_intersection_of_Int:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   158
   "\<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
   159
           \<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
   160
  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
   161
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   162
lemma arbitrary_union_of_Int_eq:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   163
  "(\<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
   164
               \<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
   165
   \<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
   166
proof
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   167
  assume ?lhs
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   168
  then show ?rhs
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   169
    by (simp add: arbitrary_union_of_inc)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   170
next
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   171
  assume R: ?rhs
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   172
  show ?lhs
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   173
  proof clarify
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   174
    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
   175
    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
   176
    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
   177
      by (auto simp: union_of_def)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   178
    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
   179
     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
   180
   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
   181
     by (simp add: Int_UN_distrib2 *)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   182
 qed
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   183
qed
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   184
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   185
lemma arbitrary_intersection_of_Un_eq:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   186
   "(\<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
   187
               \<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
   188
        (\<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
   189
  apply (simp add: arbitrary_intersection_of_complement)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   190
  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
   191
  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
   192
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   193
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
   194
  by (simp add: union_of_empty)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   195
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   196
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
   197
  by (simp add: intersection_of_empty)
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_union_of_inc:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   200
   "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
   201
  by (simp add: union_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_intersection_of_inc:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   204
   "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
   205
  by (simp add: intersection_of_inc)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   206
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   207
lemma finite_union_of_complement:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   208
  "(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
   209
  unfolding union_of_def intersection_of_def
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   210
  apply safe
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   211
   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
   212
  done
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_intersection_of_complement:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   215
   "(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
   216
  by (simp add: finite_union_of_complement)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   217
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   218
lemma finite_union_of_idempot [simp]:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   219
  "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
   220
proof -
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   221
  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
   222
  proof -
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   223
    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
   224
      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
   225
    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
   226
      by metis
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   227
    then show ?thesis
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   228
      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
   229
      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
   230
  qed
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   231
  moreover
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   232
  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
   233
    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
   234
  ultimately show ?thesis
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   235
    by force
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   236
qed
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_intersection_of_idempot [simp]:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   239
   "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
   240
  by (force simp: finite_intersection_of_complement)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   241
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   242
lemma finite_union_of_Union:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   243
   "\<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
   244
  using finite_union_of_idempot [of P]
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   245
  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
   246
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   247
lemma finite_union_of_Un:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   248
   "\<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
   249
  by (auto simp: union_of_def)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   250
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   251
lemma finite_intersection_of_Inter:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   252
   "\<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
   253
  using finite_intersection_of_idempot [of P]
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   254
  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
   255
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   256
lemma finite_intersection_of_Int:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   257
   "\<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
   258
           \<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
   259
  by (auto simp: intersection_of_def)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   260
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   261
lemma finite_union_of_Int_eq:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   262
   "(\<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
   263
    \<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
   264
(is "?lhs = ?rhs")
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   265
proof
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   266
  assume ?lhs
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   267
  then show ?rhs
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   268
    by (simp add: finite_union_of_inc)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   269
next
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   270
  assume R: ?rhs
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   271
  show ?lhs
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   272
  proof clarify
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   273
    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
   274
    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
   275
    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
   276
      by (auto simp: union_of_def)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   277
    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
   278
      using R
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   279
      by (blast intro: finite_union_of_Union)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   280
   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
   281
     by (simp add: Int_UN_distrib2 *)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   282
 qed
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   283
qed
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   284
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   285
lemma finite_intersection_of_Un_eq:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   286
   "(\<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
   287
               (finite intersection_of P) T
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   288
               \<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
   289
        (\<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
   290
  apply (simp add: finite_intersection_of_complement)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   291
  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
   292
  by (metis (no_types, lifting) double_compl)
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
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   295
abbreviation finite' :: "'a set \<Rightarrow> bool"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   296
  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
   297
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   298
lemma finite'_intersection_of_Int:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   299
   "\<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
   300
           \<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
   301
  by (auto simp: intersection_of_def)
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
lemma finite'_intersection_of_inc:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   304
   "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
   305
  by (simp add: intersection_of_inc)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   306
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   307
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   308
subsection \<open>The ``Relative to'' operator\<close>
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   309
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   310
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
   311
(open, closed, borel, fsigma, gdelta etc.)\<close>
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   312
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   313
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
   314
  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
   315
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   316
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
   317
  by (simp add: 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 relative_to_imp_subset:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   320
   "(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
   321
  by (auto simp: relative_to_def)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   322
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   323
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
   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
69918
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69325
diff changeset
   326
lemma relative_toE: "\<lbrakk>(P relative_to U) S; \<And>S. P S \<Longrightarrow> Q(U \<inter> S)\<rbrakk> \<Longrightarrow> Q S"
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69325
diff changeset
   327
  by (auto simp: relative_to_def)
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69325
diff changeset
   328
69004
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   329
lemma relative_to_inc:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   330
   "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
   331
  by (auto simp: relative_to_def)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   332
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   333
lemma relative_to_relative_to [simp]:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   334
   "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
   335
  unfolding relative_to_def
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   336
  by auto
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   337
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   338
lemma relative_to_compl:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   339
   "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
   340
  unfolding relative_to_def
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   341
  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
   342
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   343
lemma relative_to_subset:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   344
   "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
   345
  unfolding relative_to_def by auto
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   346
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   347
lemma relative_to_subset_trans:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   348
   "(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
   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_mono:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   352
   "\<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
   353
  unfolding relative_to_def by auto
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   354
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   355
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
   356
  unfolding relative_to_def by auto
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   357
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   358
lemma relative_to_Int:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   359
   "\<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
   360
         \<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
   361
  unfolding relative_to_def by auto
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   362
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   363
lemma relative_to_Un:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   364
   "\<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
   365
         \<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
   366
  unfolding relative_to_def by auto
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   367
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   368
lemma arbitrary_union_of_relative_to:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   369
  "((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
   370
proof -
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   371
  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
   372
  proof -
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   373
    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
   374
      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
   375
    then show ?thesis
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   376
      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
   377
      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
   378
  qed
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   379
  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
   380
  proof -
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   381
    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
   382
      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
   383
    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
   384
      by metis
69325
4b6ddc5989fc removed legacy input syntax
haftmann
parents: 69313
diff changeset
   385
    then have "\<exists>\<U>'\<subseteq>Collect P. \<Union>\<U>' = \<Union> (f ` \<U>)"
69004
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   386
      by (metis image_subset_iff mem_Collect_eq)
69325
4b6ddc5989fc removed legacy input syntax
haftmann
parents: 69313
diff changeset
   387
    moreover have eq: "U \<inter> \<Union> (f ` \<U>) = \<Union>\<U>"
69004
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   388
      using f by auto
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   389
    ultimately show ?thesis
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   390
      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
   391
      by metis
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   392
  qed
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   393
  ultimately show ?thesis
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   394
    by blast
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   395
qed
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   396
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   397
lemma finite_union_of_relative_to:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   398
  "((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
   399
proof -
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   400
  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
   401
  proof -
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   402
    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
   403
      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
   404
    then show ?thesis
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   405
      unfolding relative_to_def union_of_def
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   406
      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
   407
  qed
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   408
  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
   409
  proof -
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   410
    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
   411
      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
   412
    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
   413
      by metis
69325
4b6ddc5989fc removed legacy input syntax
haftmann
parents: 69313
diff changeset
   414
    then have "\<exists>\<U>'\<subseteq>Collect P. \<Union>\<U>' = \<Union> (f ` \<U>)"
69004
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   415
      by (metis image_subset_iff mem_Collect_eq)
69325
4b6ddc5989fc removed legacy input syntax
haftmann
parents: 69313
diff changeset
   416
    moreover have eq: "U \<inter> \<Union> (f ` \<U>) = \<Union>\<U>"
69004
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   417
      using f by auto
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   418
    ultimately show ?thesis
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   419
      using \<open>finite \<U>\<close> f
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   420
      unfolding relative_to_def union_of_def \<open>S = \<Union>\<U>\<close>
69325
4b6ddc5989fc removed legacy input syntax
haftmann
parents: 69313
diff changeset
   421
      by (rule_tac x="\<Union> (f ` \<U>)" in exI) (metis finite_imageI image_subsetI mem_Collect_eq)
69004
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   422
  qed
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   423
  ultimately show ?thesis
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   424
    by blast
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   425
qed
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   426
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   427
lemma countable_union_of_relative_to:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   428
  "((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
   429
proof -
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   430
  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
   431
  proof -
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   432
    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
   433
      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
   434
    then show ?thesis
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   435
      unfolding relative_to_def union_of_def
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   436
      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
   437
  qed
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   438
  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
   439
  proof -
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   440
    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
   441
      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
   442
    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
   443
      by metis
69325
4b6ddc5989fc removed legacy input syntax
haftmann
parents: 69313
diff changeset
   444
    then have "\<exists>\<U>'\<subseteq>Collect P. \<Union>\<U>' = \<Union> (f ` \<U>)"
69004
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   445
      by (metis image_subset_iff mem_Collect_eq)
69325
4b6ddc5989fc removed legacy input syntax
haftmann
parents: 69313
diff changeset
   446
    moreover have eq: "U \<inter> \<Union> (f ` \<U>) = \<Union>\<U>"
69004
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   447
      using f by auto
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   448
    ultimately show ?thesis
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   449
      using \<open>countable \<U>\<close> f
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   450
      unfolding relative_to_def union_of_def \<open>S = \<Union>\<U>\<close>
69325
4b6ddc5989fc removed legacy input syntax
haftmann
parents: 69313
diff changeset
   451
      by (rule_tac x="\<Union> (f ` \<U>)" in exI) (metis countable_image image_subsetI mem_Collect_eq)
69004
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   452
  qed
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   453
  ultimately show ?thesis
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   454
    by blast
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   455
qed
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   456
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   457
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   458
lemma arbitrary_intersection_of_relative_to:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   459
  "((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
   460
proof -
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   461
  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
   462
  proof -
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   463
    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
   464
      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
   465
    show ?thesis
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   466
      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
   467
    proof (intro exI conjI)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   468
      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
   469
        using \<U> by blast+
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   470
    qed auto
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   471
  qed
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   472
  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
   473
  proof -
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   474
    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
   475
      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
   476
    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
   477
      by metis
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 69004
diff changeset
   478
    then have "f `  \<U> \<subseteq> Collect P"
b021008c5397 removed legacy input syntax
haftmann
parents: 69004
diff changeset
   479
      by auto
b021008c5397 removed legacy input syntax
haftmann
parents: 69004
diff changeset
   480
    moreover have eq: "U \<inter> \<Inter>(f ` \<U>) = U \<inter> \<Inter>\<U>"
69004
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   481
      using f by auto
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   482
    ultimately show ?thesis
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   483
      unfolding relative_to_def intersection_of_def arbitrary_def \<open>S = U \<inter> \<Inter>\<U>\<close>
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 69004
diff changeset
   484
      by auto
69004
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   485
  qed
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   486
  ultimately show ?thesis
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   487
    by blast
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   488
qed
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   489
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   490
lemma finite_intersection_of_relative_to:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   491
  "((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
   492
proof -
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   493
  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
   494
  proof -
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   495
    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
   496
      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
   497
    show ?thesis
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   498
      unfolding relative_to_def intersection_of_def
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   499
    proof (intro exI conjI)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   500
      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
   501
        using \<U> by blast+
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   502
      show "finite ((\<inter>) U ` \<U>)"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   503
        by (simp add: \<open>finite \<U>\<close>)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   504
    qed auto
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   505
  qed
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   506
  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
   507
  proof -
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   508
    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
   509
      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
   510
    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
   511
      by metis
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 69004
diff changeset
   512
    then have "f `  \<U> \<subseteq> Collect P"
b021008c5397 removed legacy input syntax
haftmann
parents: 69004
diff changeset
   513
      by auto
b021008c5397 removed legacy input syntax
haftmann
parents: 69004
diff changeset
   514
    moreover have eq: "U \<inter> \<Inter> (f ` \<U>) = U \<inter> \<Inter> \<U>"
69004
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   515
      using f by auto
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   516
    ultimately show ?thesis
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   517
      unfolding relative_to_def intersection_of_def \<open>S = U \<inter> \<Inter>\<U>\<close>
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 69004
diff changeset
   518
      using \<open>finite \<U>\<close>
b021008c5397 removed legacy input syntax
haftmann
parents: 69004
diff changeset
   519
      by auto
69004
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   520
  qed
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   521
  ultimately show ?thesis
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   522
    by blast
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   523
qed
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   524
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   525
lemma countable_intersection_of_relative_to:
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   526
  "((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
   527
proof -
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   528
  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
   529
  proof -
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   530
    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
   531
      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
   532
    show ?thesis
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   533
      unfolding relative_to_def intersection_of_def
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   534
    proof (intro exI conjI)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   535
      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
   536
        using \<U> by blast+
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   537
      show "countable ((\<inter>) U ` \<U>)"
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   538
        by (simp add: \<open>countable \<U>\<close>)
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   539
    qed auto
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   540
  qed
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   541
  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
   542
  proof -
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   543
    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
   544
      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
   545
    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
   546
      by metis
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 69004
diff changeset
   547
    then have "f `  \<U> \<subseteq> Collect P"
b021008c5397 removed legacy input syntax
haftmann
parents: 69004
diff changeset
   548
      by auto
b021008c5397 removed legacy input syntax
haftmann
parents: 69004
diff changeset
   549
    moreover have eq: "U \<inter> \<Inter> (f ` \<U>) = U \<inter> \<Inter> \<U>"
69004
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   550
      using f by auto
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   551
    ultimately show ?thesis
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   552
      unfolding relative_to_def intersection_of_def \<open>S = U \<inter> \<Inter>\<U>\<close>
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 69004
diff changeset
   553
      using \<open>countable \<U>\<close> countable_image
b021008c5397 removed legacy input syntax
haftmann
parents: 69004
diff changeset
   554
      by auto
69004
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   555
  qed
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   556
  ultimately show ?thesis
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   557
    by blast
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   558
qed
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   559
f6a0c8115e9c Set idioms theory "finite intersection_of open", etc.
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   560
end