src/HOL/Analysis/FSigma.thy
author wenzelm
Mon, 26 Jun 2023 23:20:32 +0200
changeset 78209 50c5be88ad59
parent 77939 98879407d33c
child 78250 400aecdfd71f
permissions -rw-r--r--
tuned signature;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
77939
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     1
(*  Author:     L C Paulson, University of Cambridge [ported from HOL Light, metric.ml] *)
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     2
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     3
section \<open>$F$-Sigma and $G$-Delta sets in a Topological Space\<close>
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     4
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     5
text \<open>An $F$-sigma set is a countable union of closed sets; a $G$-delta set is the dual.\<close>
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     6
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     7
theory FSigma
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     8
  imports Abstract_Topology
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     9
begin
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    10
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    11
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    12
definition fsigma_in 
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    13
  where "fsigma_in X \<equiv> countable union_of closedin X"
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    14
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    15
definition gdelta_in 
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    16
  where "gdelta_in X \<equiv> (countable intersection_of openin X) relative_to topspace X"
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    17
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    18
lemma fsigma_in_ascending:
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    19
   "fsigma_in X S \<longleftrightarrow> (\<exists>C. (\<forall>n. closedin X (C n)) \<and> (\<forall>n. C n \<subseteq> C(Suc n)) \<and> \<Union> (range C) = S)"
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    20
  unfolding fsigma_in_def
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    21
  by (metis closedin_Un countable_union_of_ascending closedin_empty)
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    22
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    23
lemma gdelta_in_alt:
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    24
   "gdelta_in X S \<longleftrightarrow> S \<subseteq> topspace X \<and> (countable intersection_of openin X) S"
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    25
proof -
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    26
  have "(countable intersection_of openin X) (topspace X)"
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    27
    by (simp add: countable_intersection_of_inc)
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    28
  then show ?thesis
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    29
    unfolding gdelta_in_def
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    30
    by (metis countable_intersection_of_inter relative_to_def relative_to_imp_subset relative_to_subset)
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    31
qed
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    32
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    33
lemma fsigma_in_subset: "fsigma_in X S \<Longrightarrow> S \<subseteq> topspace X"
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    34
  using closedin_subset by (fastforce simp: fsigma_in_def union_of_def subset_iff)
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    35
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    36
lemma gdelta_in_subset: "gdelta_in X S \<Longrightarrow> S \<subseteq> topspace X"
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    37
  by (simp add: gdelta_in_alt)
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    38
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    39
lemma closed_imp_fsigma_in: "closedin X S \<Longrightarrow> fsigma_in X S"
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    40
  by (simp add: countable_union_of_inc fsigma_in_def)
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    41
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    42
lemma open_imp_gdelta_in: "openin X S \<Longrightarrow> gdelta_in X S"
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    43
  by (simp add: countable_intersection_of_inc gdelta_in_alt openin_subset)
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    44
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    45
lemma fsigma_in_empty [iff]: "fsigma_in X {}"
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    46
  by (simp add: closed_imp_fsigma_in)
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    47
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    48
lemma gdelta_in_empty [iff]: "gdelta_in X {}"
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    49
  by (simp add: open_imp_gdelta_in)
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    50
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    51
lemma fsigma_in_topspace [iff]: "fsigma_in X (topspace X)"
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    52
  by (simp add: closed_imp_fsigma_in)
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    53
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    54
lemma gdelta_in_topspace [iff]: "gdelta_in X (topspace X)"
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    55
  by (simp add: open_imp_gdelta_in)
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    56
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    57
lemma fsigma_in_Union:
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    58
   "\<lbrakk>countable T; \<And>S. S \<in> T \<Longrightarrow> fsigma_in X S\<rbrakk> \<Longrightarrow> fsigma_in X (\<Union> T)"
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    59
  by (simp add: countable_union_of_Union fsigma_in_def)
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    60
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    61
lemma fsigma_in_Un:
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    62
   "\<lbrakk>fsigma_in X S; fsigma_in X T\<rbrakk> \<Longrightarrow> fsigma_in X (S \<union> T)"
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    63
  by (simp add: countable_union_of_Un fsigma_in_def)
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    64
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    65
lemma fsigma_in_Int:
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    66
   "\<lbrakk>fsigma_in X S; fsigma_in X T\<rbrakk> \<Longrightarrow> fsigma_in X (S \<inter> T)"
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    67
  by (simp add: closedin_Int countable_union_of_Int fsigma_in_def)
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    68
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    69
lemma gdelta_in_Inter:
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    70
   "\<lbrakk>countable T; T\<noteq>{}; \<And>S. S \<in> T \<Longrightarrow> gdelta_in X S\<rbrakk> \<Longrightarrow> gdelta_in X (\<Inter> T)"
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    71
  by (simp add: Inf_less_eq countable_intersection_of_Inter gdelta_in_alt)
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    72
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    73
lemma gdelta_in_Int:
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    74
   "\<lbrakk>gdelta_in X S; gdelta_in X T\<rbrakk> \<Longrightarrow> gdelta_in X (S \<inter> T)"
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    75
  by (simp add: countable_intersection_of_inter gdelta_in_alt le_infI2)
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    76
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    77
lemma gdelta_in_Un:
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    78
   "\<lbrakk>gdelta_in X S; gdelta_in X T\<rbrakk> \<Longrightarrow> gdelta_in X (S \<union> T)"
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    79
  by (simp add: countable_intersection_of_union gdelta_in_alt openin_Un)
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    80
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    81
lemma fsigma_in_diff:
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    82
  assumes "fsigma_in X S" "gdelta_in X T"
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    83
  shows "fsigma_in X (S - T)"
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    84
proof -
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    85
  have [simp]: "S - (S \<inter> T) = S - T" for S T :: "'a set"
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    86
    by auto
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    87
  have [simp]: "topspace X - \<Inter>\<T> = (\<Union>T\<in>\<T>. topspace X - T)" for \<T>
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    88
    by auto
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    89
  have "\<And>\<T>. \<lbrakk>countable \<T>; \<T> \<subseteq> Collect (openin X)\<rbrakk> \<Longrightarrow>
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    90
             (countable union_of closedin X) (\<Union> ((-) (topspace X) ` \<T>))"
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    91
    by (metis Ball_Collect countable_union_of_UN countable_union_of_inc openin_closedin_eq)
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    92
  then have "\<forall>S. gdelta_in X S \<longrightarrow> fsigma_in X (topspace X - S)"
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    93
    by (simp add: fsigma_in_def gdelta_in_def all_relative_to all_intersection_of del: UN_simps)
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    94
  then show ?thesis
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    95
    by (metis Diff_Int2 Diff_Int_distrib2 assms fsigma_in_Int fsigma_in_subset inf.absorb_iff2)
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    96
qed
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    97
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    98
lemma gdelta_in_diff:
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    99
  assumes "gdelta_in X S" "fsigma_in X T"
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   100
  shows "gdelta_in X (S - T)"
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   101
proof -
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   102
  have [simp]: "topspace X - \<Union>\<T> = topspace X \<inter> (\<Inter>T\<in>\<T>. topspace X - T)" for \<T>
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   103
    by auto
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   104
  have "\<And>\<T>. \<lbrakk>countable \<T>; \<T> \<subseteq> Collect (closedin X)\<rbrakk>
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   105
         \<Longrightarrow> (countable intersection_of openin X relative_to topspace X)
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   106
              (topspace X \<inter> \<Inter> ((-) (topspace X) ` \<T>))"
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   107
    by (metis Ball_Collect closedin_def countable_intersection_of_INT countable_intersection_of_inc relative_to_inc)
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   108
  then have "\<forall>S. fsigma_in X S \<longrightarrow> gdelta_in X (topspace X - S)"
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   109
    by (simp add: fsigma_in_def gdelta_in_def all_union_of del: INT_simps)
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   110
  then show ?thesis
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   111
    by (metis Diff_Int2 Diff_Int_distrib2 assms gdelta_in_Int gdelta_in_alt inf.orderE inf_commute)
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   112
qed
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   113
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   114
lemma gdelta_in_fsigma_in:
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   115
   "gdelta_in X S \<longleftrightarrow> S \<subseteq> topspace X \<and> fsigma_in X (topspace X - S)"
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   116
  by (metis double_diff fsigma_in_diff fsigma_in_topspace gdelta_in_alt gdelta_in_diff gdelta_in_topspace)
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   117
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   118
lemma fsigma_in_gdelta_in:
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   119
   "fsigma_in X S \<longleftrightarrow> S \<subseteq> topspace X \<and> gdelta_in X (topspace X - S)"
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   120
  by (metis Diff_Diff_Int fsigma_in_subset gdelta_in_fsigma_in inf.absorb_iff2)
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   121
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   122
lemma gdelta_in_descending:
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   123
   "gdelta_in X S \<longleftrightarrow> (\<exists>C. (\<forall>n. openin X (C n)) \<and> (\<forall>n. C(Suc n) \<subseteq> C n) \<and> \<Inter>(range C) = S)" (is "?lhs=?rhs")
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   124
proof
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   125
  assume ?lhs
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   126
  then obtain C where C: "S \<subseteq> topspace X" "\<And>n. closedin X (C n)" 
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   127
                         "\<And>n. C n \<subseteq> C(Suc n)" "\<Union>(range C) = topspace X - S"
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   128
    by (meson fsigma_in_ascending gdelta_in_fsigma_in)
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   129
  define D where "D \<equiv> \<lambda>n. topspace X - C n"
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   130
  have "\<And>n. openin X (D n) \<and> D (Suc n) \<subseteq> D n"
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   131
    by (simp add: Diff_mono C D_def openin_diff)
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   132
  moreover have "\<Inter>(range D) = S"
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   133
    by (simp add: Diff_Diff_Int Int_absorb1 C D_def)
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   134
  ultimately show ?rhs
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   135
    by metis
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   136
next
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   137
  assume ?rhs
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   138
  then obtain C where "S \<subseteq> topspace X" 
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   139
                and C: "\<And>n. openin X (C n)" "\<And>n. C(Suc n) \<subseteq> C n" "\<Inter>(range C) = S"
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   140
    using openin_subset by fastforce
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   141
  define D where "D \<equiv> \<lambda>n. topspace X - C n"
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   142
  have "\<And>n. closedin X (D n) \<and> D n \<subseteq> D(Suc n)"
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   143
    by (simp add: Diff_mono C closedin_diff D_def)
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   144
  moreover have "\<Union>(range D) = topspace X - S"
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   145
    using C D_def by blast
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   146
  ultimately show ?lhs
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   147
    by (metis \<open>S \<subseteq> topspace X\<close> fsigma_in_ascending gdelta_in_fsigma_in)
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   148
qed
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   149
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   150
lemma homeomorphic_map_fsigmaness:
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   151
  assumes f: "homeomorphic_map X Y f" and "U \<subseteq> topspace X"
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   152
  shows "fsigma_in Y (f ` U) \<longleftrightarrow> fsigma_in X U"  (is "?lhs=?rhs")
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   153
proof -
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   154
  obtain g where g: "homeomorphic_maps X Y f g" and g: "homeomorphic_map Y X g"
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   155
    and 1: "(\<forall>x \<in> topspace X. g(f x) = x)" and 2: "(\<forall>y \<in> topspace Y. f(g y) = y)"
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   156
    using assms homeomorphic_map_maps by (metis homeomorphic_maps_map)
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   157
  show ?thesis
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   158
  proof
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   159
    assume ?lhs
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   160
    then obtain \<V> where "countable \<V>" and \<V>: "\<V> \<subseteq> Collect (closedin Y)" "\<Union>\<V> = f`U"
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   161
      by (force simp: fsigma_in_def union_of_def)
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   162
    define \<U> where "\<U> \<equiv> image (image g) \<V>"
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   163
    have "countable \<U>"
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   164
      using \<U>_def \<open>countable \<V>\<close> by blast
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   165
    moreover have "\<U> \<subseteq> Collect (closedin X)"
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   166
      using f g homeomorphic_map_closedness_eq \<V> unfolding \<U>_def by blast
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   167
    moreover have "\<Union>\<U> \<subseteq> U"
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   168
      unfolding \<U>_def  by (smt (verit) assms 1 \<V> image_Union image_iff in_mono subsetI)
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   169
    moreover have "U \<subseteq> \<Union>\<U>"
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   170
      unfolding \<U>_def using assms \<V>
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   171
      by (smt (verit, del_insts) "1" imageI image_Union subset_iff)
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   172
    ultimately show ?rhs
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   173
      by (metis fsigma_in_def subset_antisym union_of_def)
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   174
  next
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   175
    assume ?rhs
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   176
    then obtain \<V> where "countable \<V>" and \<V>: "\<V> \<subseteq> Collect (closedin X)" "\<Union>\<V> = U"
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   177
      by (auto simp: fsigma_in_def union_of_def)
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   178
    define \<U> where "\<U> \<equiv> image (image f) \<V>"
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   179
    have "countable \<U>"
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   180
      using \<U>_def \<open>countable \<V>\<close> by blast
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   181
    moreover have "\<U> \<subseteq> Collect (closedin Y)"
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   182
      using f g homeomorphic_map_closedness_eq \<V> unfolding \<U>_def by blast
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   183
    moreover have "\<Union>\<U> = f`U"
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   184
      unfolding \<U>_def using \<V> by blast
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   185
    ultimately show ?lhs
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   186
      by (metis fsigma_in_def union_of_def)
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   187
  qed
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   188
qed
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   189
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   190
lemma homeomorphic_map_fsigmaness_eq:
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   191
   "homeomorphic_map X Y f
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   192
        \<Longrightarrow> (fsigma_in X U \<longleftrightarrow> U \<subseteq> topspace X \<and> fsigma_in Y (f ` U))"
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   193
  by (metis fsigma_in_subset homeomorphic_map_fsigmaness)
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   194
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   195
lemma homeomorphic_map_gdeltaness:
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   196
  assumes "homeomorphic_map X Y f" "U \<subseteq> topspace X"
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   197
  shows "gdelta_in Y (f ` U) \<longleftrightarrow> gdelta_in X U"
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   198
proof -
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   199
  have "topspace Y - f ` U = f ` (topspace X - U)"
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   200
    by (metis (no_types, lifting) Diff_subset assms homeomorphic_eq_everything_map inj_on_image_set_diff)
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   201
  then show ?thesis
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   202
    using assms homeomorphic_imp_surjective_map
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   203
    by (fastforce simp: gdelta_in_fsigma_in homeomorphic_map_fsigmaness_eq)
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   204
qed
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   205
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   206
lemma homeomorphic_map_gdeltaness_eq:
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   207
   "homeomorphic_map X Y f
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   208
    \<Longrightarrow> gdelta_in X U \<longleftrightarrow> U \<subseteq> topspace X \<and> gdelta_in Y (f ` U)"
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   209
  by (meson gdelta_in_subset homeomorphic_map_gdeltaness)
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   210
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   211
lemma fsigma_in_relative_to:
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   212
   "(fsigma_in X relative_to S) = fsigma_in (subtopology X S)"
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   213
  by (simp add: fsigma_in_def countable_union_of_relative_to closedin_relative_to)
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   214
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   215
lemma fsigma_in_subtopology:
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   216
   "fsigma_in (subtopology X U) S \<longleftrightarrow> (\<exists>T. fsigma_in X T \<and> S = T \<inter> U)"
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   217
  by (metis fsigma_in_relative_to inf_commute relative_to_def)
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   218
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   219
lemma gdelta_in_relative_to:
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   220
   "(gdelta_in X relative_to S) = gdelta_in (subtopology X S)"
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   221
  apply (simp add: gdelta_in_def)
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   222
  by (metis countable_intersection_of_relative_to openin_relative_to subtopology_restrict)
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   223
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   224
lemma gdelta_in_subtopology:
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   225
   "gdelta_in (subtopology X U) S \<longleftrightarrow> (\<exists>T. gdelta_in X T \<and> S = T \<inter> U)"
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   226
  by (metis gdelta_in_relative_to inf_commute relative_to_def)
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   227
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   228
lemma fsigma_in_fsigma_subtopology:
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   229
   "fsigma_in X S \<Longrightarrow> (fsigma_in (subtopology X S) T \<longleftrightarrow> fsigma_in X T \<and> T \<subseteq> S)"
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   230
  by (metis fsigma_in_Int fsigma_in_gdelta_in fsigma_in_subtopology inf.orderE topspace_subtopology_subset)
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   231
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   232
lemma gdelta_in_gdelta_subtopology:
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   233
   "gdelta_in X S \<Longrightarrow> (gdelta_in (subtopology X S) T \<longleftrightarrow> gdelta_in X T \<and> T \<subseteq> S)"
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   234
  by (metis gdelta_in_Int gdelta_in_subset gdelta_in_subtopology inf.orderE topspace_subtopology_subset)
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   235
98879407d33c Two new theories containing material ported from HOL Light about abstract topology
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   236
end