src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy
author wenzelm
Fri, 01 Feb 2013 21:58:00 +0100
changeset 51074 f95817852bdd
parent 50973 4a2c82644889
child 50998 501200635659
permissions -rw-r--r--
provide jdk-7u13 (still inactive, potentially pointless);
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
33714
eb2574ac4173 Added new lemmas to Euclidean Space by Robert Himmelmann
hoelzl
parents: 33324
diff changeset
     1
(*  title:      HOL/Library/Topology_Euclidian_Space.thy
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
     2
    Author:     Amine Chaieb, University of Cambridge
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
     3
    Author:     Robert Himmelmann, TU Muenchen
44075
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
     4
    Author:     Brian Huffman, Portland State University
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
     5
*)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
     6
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
     7
header {* Elementary topology in Euclidean space. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
     8
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
     9
theory Topology_Euclidean_Space
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    10
imports
50938
hoelzl
parents: 50937
diff changeset
    11
  Complex_Main
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
    12
  "~~/src/HOL/Library/Countable_Set"
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    13
  "~~/src/HOL/Library/Glbs"
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
    14
  "~~/src/HOL/Library/FuncSet"
50938
hoelzl
parents: 50937
diff changeset
    15
  Linear_Algebra
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    16
  Norm_Arith
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    17
begin
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    18
50972
d2c6a0a7fcdf tuned proof
hoelzl
parents: 50971
diff changeset
    19
lemma dist_0_norm:
d2c6a0a7fcdf tuned proof
hoelzl
parents: 50971
diff changeset
    20
  fixes x :: "'a::real_normed_vector"
d2c6a0a7fcdf tuned proof
hoelzl
parents: 50971
diff changeset
    21
  shows "dist 0 x = norm x"
d2c6a0a7fcdf tuned proof
hoelzl
parents: 50971
diff changeset
    22
unfolding dist_norm by simp
d2c6a0a7fcdf tuned proof
hoelzl
parents: 50971
diff changeset
    23
50943
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
    24
lemma dist_double: "dist x y < d / 2 \<Longrightarrow> dist x z < d / 2 \<Longrightarrow> dist y z < d"
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
    25
  using dist_triangle[of y z x] by (simp add: dist_commute)
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
    26
50972
d2c6a0a7fcdf tuned proof
hoelzl
parents: 50971
diff changeset
    27
(* LEGACY *)
d2c6a0a7fcdf tuned proof
hoelzl
parents: 50971
diff changeset
    28
lemma lim_subseq: "subseq r \<Longrightarrow> s ----> l \<Longrightarrow> (s o r) ----> l"
d2c6a0a7fcdf tuned proof
hoelzl
parents: 50971
diff changeset
    29
  by (rule LIMSEQ_subseq_LIMSEQ)
d2c6a0a7fcdf tuned proof
hoelzl
parents: 50971
diff changeset
    30
50942
1aa1a7991fd9 move auxiliary lemma to top
hoelzl
parents: 50941
diff changeset
    31
(* TODO: Move this to RComplete.thy -- would need to include Glb into RComplete *)
1aa1a7991fd9 move auxiliary lemma to top
hoelzl
parents: 50941
diff changeset
    32
lemma real_isGlb_unique: "[| isGlb R S x; isGlb R S y |] ==> x = (y::real)"
1aa1a7991fd9 move auxiliary lemma to top
hoelzl
parents: 50941
diff changeset
    33
  apply (frule isGlb_isLb)
1aa1a7991fd9 move auxiliary lemma to top
hoelzl
parents: 50941
diff changeset
    34
  apply (frule_tac x = y in isGlb_isLb)
1aa1a7991fd9 move auxiliary lemma to top
hoelzl
parents: 50941
diff changeset
    35
  apply (blast intro!: order_antisym dest!: isGlb_le_isLb)
1aa1a7991fd9 move auxiliary lemma to top
hoelzl
parents: 50941
diff changeset
    36
  done
1aa1a7991fd9 move auxiliary lemma to top
hoelzl
parents: 50941
diff changeset
    37
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
    38
lemma countable_PiE: 
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
    39
  "finite I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> countable (F i)) \<Longrightarrow> countable (PiE I F)"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
    40
  by (induct I arbitrary: F rule: finite_induct) (auto simp: PiE_insert_eq)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
    41
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    42
subsection {* Topological Basis *}
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    43
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    44
context topological_space
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    45
begin
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    46
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    47
definition "topological_basis B =
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    48
  ((\<forall>b\<in>B. open b) \<and> (\<forall>x. open x \<longrightarrow> (\<exists>B'. B' \<subseteq> B \<and> Union B' = x)))"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    49
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    50
lemma topological_basis_iff:
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    51
  assumes "\<And>B'. B' \<in> B \<Longrightarrow> open B'"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    52
  shows "topological_basis B \<longleftrightarrow> (\<forall>O'. open O' \<longrightarrow> (\<forall>x\<in>O'. \<exists>B'\<in>B. x \<in> B' \<and> B' \<subseteq> O'))"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    53
    (is "_ \<longleftrightarrow> ?rhs")
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    54
proof safe
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    55
  fix O' and x::'a
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    56
  assume H: "topological_basis B" "open O'" "x \<in> O'"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    57
  hence "(\<exists>B'\<subseteq>B. \<Union>B' = O')" by (simp add: topological_basis_def)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    58
  then obtain B' where "B' \<subseteq> B" "O' = \<Union>B'" by auto
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    59
  thus "\<exists>B'\<in>B. x \<in> B' \<and> B' \<subseteq> O'" using H by auto
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    60
next
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    61
  assume H: ?rhs
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    62
  show "topological_basis B" using assms unfolding topological_basis_def
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    63
  proof safe
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    64
    fix O'::"'a set" assume "open O'"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    65
    with H obtain f where "\<forall>x\<in>O'. f x \<in> B \<and> x \<in> f x \<and> f x \<subseteq> O'"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    66
      by (force intro: bchoice simp: Bex_def)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    67
    thus "\<exists>B'\<subseteq>B. \<Union>B' = O'"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    68
      by (auto intro: exI[where x="{f x |x. x \<in> O'}"])
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    69
  qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    70
qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    71
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    72
lemma topological_basisI:
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    73
  assumes "\<And>B'. B' \<in> B \<Longrightarrow> open B'"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    74
  assumes "\<And>O' x. open O' \<Longrightarrow> x \<in> O' \<Longrightarrow> \<exists>B'\<in>B. x \<in> B' \<and> B' \<subseteq> O'"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    75
  shows "topological_basis B"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    76
  using assms by (subst topological_basis_iff) auto
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    77
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    78
lemma topological_basisE:
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    79
  fixes O'
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    80
  assumes "topological_basis B"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    81
  assumes "open O'"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    82
  assumes "x \<in> O'"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    83
  obtains B' where "B' \<in> B" "x \<in> B'" "B' \<subseteq> O'"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    84
proof atomize_elim
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    85
  from assms have "\<And>B'. B'\<in>B \<Longrightarrow> open B'" by (simp add: topological_basis_def)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    86
  with topological_basis_iff assms
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    87
  show  "\<exists>B'. B' \<in> B \<and> x \<in> B' \<and> B' \<subseteq> O'" using assms by (simp add: Bex_def)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    88
qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    89
50094
84ddcf5364b4 allow arbitrary enumerations of basis in locale for generation of borel sets
immler
parents: 50087
diff changeset
    90
lemma topological_basis_open:
84ddcf5364b4 allow arbitrary enumerations of basis in locale for generation of borel sets
immler
parents: 50087
diff changeset
    91
  assumes "topological_basis B"
84ddcf5364b4 allow arbitrary enumerations of basis in locale for generation of borel sets
immler
parents: 50087
diff changeset
    92
  assumes "X \<in> B"
84ddcf5364b4 allow arbitrary enumerations of basis in locale for generation of borel sets
immler
parents: 50087
diff changeset
    93
  shows "open X"
84ddcf5364b4 allow arbitrary enumerations of basis in locale for generation of borel sets
immler
parents: 50087
diff changeset
    94
  using assms
84ddcf5364b4 allow arbitrary enumerations of basis in locale for generation of borel sets
immler
parents: 50087
diff changeset
    95
  by (simp add: topological_basis_def)
84ddcf5364b4 allow arbitrary enumerations of basis in locale for generation of borel sets
immler
parents: 50087
diff changeset
    96
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
    97
lemma basis_dense:
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
    98
  fixes B::"'a set set" and f::"'a set \<Rightarrow> 'a"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
    99
  assumes "topological_basis B"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   100
  assumes choosefrom_basis: "\<And>B'. B' \<noteq> {} \<Longrightarrow> f B' \<in> B'"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   101
  shows "(\<forall>X. open X \<longrightarrow> X \<noteq> {} \<longrightarrow> (\<exists>B' \<in> B. f B' \<in> X))"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   102
proof (intro allI impI)
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   103
  fix X::"'a set" assume "open X" "X \<noteq> {}"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   104
  from topological_basisE[OF `topological_basis B` `open X` choosefrom_basis[OF `X \<noteq> {}`]]
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   105
  guess B' . note B' = this
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   106
  thus "\<exists>B'\<in>B. f B' \<in> X" by (auto intro!: choosefrom_basis)
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   107
qed
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   108
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   109
end
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   110
50882
a382bf90867e move prod instantiation of second_countable_topology to its definition
hoelzl
parents: 50881
diff changeset
   111
lemma topological_basis_prod:
a382bf90867e move prod instantiation of second_countable_topology to its definition
hoelzl
parents: 50881
diff changeset
   112
  assumes A: "topological_basis A" and B: "topological_basis B"
a382bf90867e move prod instantiation of second_countable_topology to its definition
hoelzl
parents: 50881
diff changeset
   113
  shows "topological_basis ((\<lambda>(a, b). a \<times> b) ` (A \<times> B))"
a382bf90867e move prod instantiation of second_countable_topology to its definition
hoelzl
parents: 50881
diff changeset
   114
  unfolding topological_basis_def
a382bf90867e move prod instantiation of second_countable_topology to its definition
hoelzl
parents: 50881
diff changeset
   115
proof (safe, simp_all del: ex_simps add: subset_image_iff ex_simps(1)[symmetric])
a382bf90867e move prod instantiation of second_countable_topology to its definition
hoelzl
parents: 50881
diff changeset
   116
  fix S :: "('a \<times> 'b) set" assume "open S"
a382bf90867e move prod instantiation of second_countable_topology to its definition
hoelzl
parents: 50881
diff changeset
   117
  then show "\<exists>X\<subseteq>A \<times> B. (\<Union>(a,b)\<in>X. a \<times> b) = S"
a382bf90867e move prod instantiation of second_countable_topology to its definition
hoelzl
parents: 50881
diff changeset
   118
  proof (safe intro!: exI[of _ "{x\<in>A \<times> B. fst x \<times> snd x \<subseteq> S}"])
a382bf90867e move prod instantiation of second_countable_topology to its definition
hoelzl
parents: 50881
diff changeset
   119
    fix x y assume "(x, y) \<in> S"
a382bf90867e move prod instantiation of second_countable_topology to its definition
hoelzl
parents: 50881
diff changeset
   120
    from open_prod_elim[OF `open S` this]
a382bf90867e move prod instantiation of second_countable_topology to its definition
hoelzl
parents: 50881
diff changeset
   121
    obtain a b where a: "open a""x \<in> a" and b: "open b" "y \<in> b" and "a \<times> b \<subseteq> S"
a382bf90867e move prod instantiation of second_countable_topology to its definition
hoelzl
parents: 50881
diff changeset
   122
      by (metis mem_Sigma_iff)
a382bf90867e move prod instantiation of second_countable_topology to its definition
hoelzl
parents: 50881
diff changeset
   123
    moreover from topological_basisE[OF A a] guess A0 .
a382bf90867e move prod instantiation of second_countable_topology to its definition
hoelzl
parents: 50881
diff changeset
   124
    moreover from topological_basisE[OF B b] guess B0 .
a382bf90867e move prod instantiation of second_countable_topology to its definition
hoelzl
parents: 50881
diff changeset
   125
    ultimately show "(x, y) \<in> (\<Union>(a, b)\<in>{X \<in> A \<times> B. fst X \<times> snd X \<subseteq> S}. a \<times> b)"
a382bf90867e move prod instantiation of second_countable_topology to its definition
hoelzl
parents: 50881
diff changeset
   126
      by (intro UN_I[of "(A0, B0)"]) auto
a382bf90867e move prod instantiation of second_countable_topology to its definition
hoelzl
parents: 50881
diff changeset
   127
  qed auto
a382bf90867e move prod instantiation of second_countable_topology to its definition
hoelzl
parents: 50881
diff changeset
   128
qed (metis A B topological_basis_open open_Times)
a382bf90867e move prod instantiation of second_countable_topology to its definition
hoelzl
parents: 50881
diff changeset
   129
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   130
subsection {* Countable Basis *}
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   131
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   132
locale countable_basis =
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   133
  fixes B::"'a::topological_space set set"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   134
  assumes is_basis: "topological_basis B"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   135
  assumes countable_basis: "countable B"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   136
begin
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   137
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   138
lemma open_countable_basis_ex:
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   139
  assumes "open X"
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   140
  shows "\<exists>B' \<subseteq> B. X = Union B'"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   141
  using assms countable_basis is_basis unfolding topological_basis_def by blast
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   142
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   143
lemma open_countable_basisE:
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   144
  assumes "open X"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   145
  obtains B' where "B' \<subseteq> B" "X = Union B'"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   146
  using assms open_countable_basis_ex by (atomize_elim) simp
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   147
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   148
lemma countable_dense_exists:
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   149
  shows "\<exists>D::'a set. countable D \<and> (\<forall>X. open X \<longrightarrow> X \<noteq> {} \<longrightarrow> (\<exists>d \<in> D. d \<in> X))"
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   150
proof -
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   151
  let ?f = "(\<lambda>B'. SOME x. x \<in> B')"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   152
  have "countable (?f ` B)" using countable_basis by simp
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   153
  with basis_dense[OF is_basis, of ?f] show ?thesis
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   154
    by (intro exI[where x="?f ` B"]) (metis (mono_tags) all_not_in_conv imageI someI)
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   155
qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   156
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   157
lemma countable_dense_setE:
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   158
  obtains D :: "'a set"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   159
  where "countable D" "\<And>X. open X \<Longrightarrow> X \<noteq> {} \<Longrightarrow> \<exists>d \<in> D. d \<in> X"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   160
  using countable_dense_exists by blast
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   161
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   162
text {* Construction of an increasing sequence approximating open sets,
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   163
  therefore basis which is closed under union. *}
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   164
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   165
definition union_closed_basis::"'a set set" where
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   166
  "union_closed_basis = (\<lambda>l. \<Union>set l) ` lists B"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   167
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   168
lemma basis_union_closed_basis: "topological_basis union_closed_basis"
50094
84ddcf5364b4 allow arbitrary enumerations of basis in locale for generation of borel sets
immler
parents: 50087
diff changeset
   169
proof (rule topological_basisI)
84ddcf5364b4 allow arbitrary enumerations of basis in locale for generation of borel sets
immler
parents: 50087
diff changeset
   170
  fix O' and x::'a assume "open O'" "x \<in> O'"
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   171
  from topological_basisE[OF is_basis this] guess B' . note B' = this
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   172
  thus "\<exists>B'\<in>union_closed_basis. x \<in> B' \<and> B' \<subseteq> O'" unfolding union_closed_basis_def
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   173
    by (auto intro!: bexI[where x="[B']"])
50094
84ddcf5364b4 allow arbitrary enumerations of basis in locale for generation of borel sets
immler
parents: 50087
diff changeset
   174
next
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   175
  fix B' assume "B' \<in> union_closed_basis"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   176
  thus "open B'"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   177
    using topological_basis_open[OF is_basis]
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   178
    by (auto simp: union_closed_basis_def)
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   179
qed
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   180
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   181
lemma countable_union_closed_basis: "countable union_closed_basis"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   182
  unfolding union_closed_basis_def using countable_basis by simp
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   183
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   184
lemmas open_union_closed_basis = topological_basis_open[OF basis_union_closed_basis]
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   185
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   186
lemma union_closed_basis_ex:
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   187
 assumes X: "X \<in> union_closed_basis"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   188
 shows "\<exists>B'. finite B' \<and> X = \<Union>B' \<and> B' \<subseteq> B"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   189
proof -
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   190
  from X obtain l where "\<And>x. x\<in>set l \<Longrightarrow> x\<in>B" "X = \<Union>set l" by (auto simp: union_closed_basis_def)
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   191
  thus ?thesis by auto
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   192
qed
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   193
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   194
lemma union_closed_basisE:
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   195
  assumes "X \<in> union_closed_basis"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   196
  obtains B' where "finite B'" "X = \<Union>B'" "B' \<subseteq> B" using union_closed_basis_ex[OF assms] by blast
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   197
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   198
lemma union_closed_basisI:
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   199
  assumes "finite B'" "X = \<Union>B'" "B' \<subseteq> B"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   200
  shows "X \<in> union_closed_basis"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   201
proof -
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   202
  from finite_list[OF `finite B'`] guess l ..
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   203
  thus ?thesis using assms unfolding union_closed_basis_def by (auto intro!: image_eqI[where x=l])
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   204
qed
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   205
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   206
lemma empty_basisI[intro]: "{} \<in> union_closed_basis"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   207
  by (rule union_closed_basisI[of "{}"]) auto
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   208
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   209
lemma union_basisI[intro]:
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   210
  assumes "X \<in> union_closed_basis" "Y \<in> union_closed_basis"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   211
  shows "X \<union> Y \<in> union_closed_basis"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   212
  using assms by (auto intro: union_closed_basisI elim!:union_closed_basisE)
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   213
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   214
lemma open_imp_Union_of_incseq:
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   215
  assumes "open X"
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   216
  shows "\<exists>S. incseq S \<and> (\<Union>j. S j) = X \<and> range S \<subseteq> union_closed_basis"
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   217
proof -
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   218
  from open_countable_basis_ex[OF `open X`]
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   219
  obtain B' where B': "B'\<subseteq>B" "X = \<Union>B'" by auto
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   220
  from this(1) countable_basis have "countable B'" by (rule countable_subset)
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   221
  show ?thesis
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   222
  proof cases
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   223
    assume "B' \<noteq> {}"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   224
    def S \<equiv> "\<lambda>n. \<Union>i\<in>{0..n}. from_nat_into B' i"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   225
    have S:"\<And>n. S n = \<Union>{from_nat_into B' i|i. i\<in>{0..n}}" unfolding S_def by force
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   226
    have "incseq S" by (force simp: S_def incseq_Suc_iff)
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   227
    moreover
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   228
    have "(\<Union>j. S j) = X" unfolding B'
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   229
    proof safe
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   230
      fix x X assume "X \<in> B'" "x \<in> X"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   231
      then obtain n where "X = from_nat_into B' n"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   232
        by (metis `countable B'` from_nat_into_surj)
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   233
      also have "\<dots> \<subseteq> S n" by (auto simp: S_def)
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   234
      finally show "x \<in> (\<Union>j. S j)" using `x \<in> X` by auto
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   235
    next
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   236
      fix x n
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   237
      assume "x \<in> S n"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   238
      also have "\<dots> = (\<Union>i\<in>{0..n}. from_nat_into B' i)"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   239
        by (simp add: S_def)
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   240
      also have "\<dots> \<subseteq> (\<Union>i. from_nat_into B' i)" by auto
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   241
      also have "\<dots> \<subseteq> \<Union>B'" using `B' \<noteq> {}` by (auto intro: from_nat_into)
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   242
      finally show "x \<in> \<Union>B'" .
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   243
    qed
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   244
    moreover have "range S \<subseteq> union_closed_basis" using B'
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   245
      by (auto intro!: union_closed_basisI[OF _ S] simp: from_nat_into `B' \<noteq> {}`)
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   246
    ultimately show ?thesis by auto
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   247
  qed (auto simp: B')
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   248
qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   249
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   250
lemma open_incseqE:
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   251
  assumes "open X"
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   252
  obtains S where "incseq S" "(\<Union>j. S j) = X" "range S \<subseteq> union_closed_basis"
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   253
  using open_imp_Union_of_incseq assms by atomize_elim
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   254
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   255
end
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   256
50883
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   257
class first_countable_topology = topological_space +
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   258
  assumes first_countable_basis:
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   259
    "\<exists>A. countable A \<and> (\<forall>a\<in>A. x \<in> a \<and> open a) \<and> (\<forall>S. open S \<and> x \<in> S \<longrightarrow> (\<exists>a\<in>A. a \<subseteq> S))"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   260
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   261
lemma (in first_countable_topology) countable_basis_at_decseq:
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   262
  obtains A :: "nat \<Rightarrow> 'a set" where
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   263
    "\<And>i. open (A i)" "\<And>i. x \<in> (A i)"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   264
    "\<And>S. open S \<Longrightarrow> x \<in> S \<Longrightarrow> eventually (\<lambda>i. A i \<subseteq> S) sequentially"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   265
proof atomize_elim
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   266
  from first_countable_basis[of x] obtain A
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   267
    where "countable A"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   268
    and nhds: "\<And>a. a \<in> A \<Longrightarrow> open a" "\<And>a. a \<in> A \<Longrightarrow> x \<in> a"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   269
    and incl: "\<And>S. open S \<Longrightarrow> x \<in> S \<Longrightarrow> \<exists>a\<in>A. a \<subseteq> S"  by auto
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   270
  then have "A \<noteq> {}" by auto
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   271
  with `countable A` have r: "A = range (from_nat_into A)" by auto
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   272
  def F \<equiv> "\<lambda>n. \<Inter>i\<le>n. from_nat_into A i"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   273
  show "\<exists>A. (\<forall>i. open (A i)) \<and> (\<forall>i. x \<in> A i) \<and>
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   274
      (\<forall>S. open S \<longrightarrow> x \<in> S \<longrightarrow> eventually (\<lambda>i. A i \<subseteq> S) sequentially)"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   275
  proof (safe intro!: exI[of _ F])
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   276
    fix i
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   277
    show "open (F i)" using nhds(1) r by (auto simp: F_def intro!: open_INT)
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   278
    show "x \<in> F i" using nhds(2) r by (auto simp: F_def)
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   279
  next
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   280
    fix S assume "open S" "x \<in> S"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   281
    from incl[OF this] obtain i where "F i \<subseteq> S"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   282
      by (subst (asm) r) (auto simp: F_def)
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   283
    moreover have "\<And>j. i \<le> j \<Longrightarrow> F j \<subseteq> F i"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   284
      by (auto simp: F_def)
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   285
    ultimately show "eventually (\<lambda>i. F i \<subseteq> S) sequentially"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   286
      by (auto simp: eventually_sequentially)
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   287
  qed
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   288
qed
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   289
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   290
lemma (in first_countable_topology) first_countable_basisE:
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   291
  obtains A where "countable A" "\<And>a. a \<in> A \<Longrightarrow> x \<in> a" "\<And>a. a \<in> A \<Longrightarrow> open a"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   292
    "\<And>S. open S \<Longrightarrow> x \<in> S \<Longrightarrow> (\<exists>a\<in>A. a \<subseteq> S)"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   293
  using first_countable_basis[of x]
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   294
  by atomize_elim auto
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   295
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   296
instance prod :: (first_countable_topology, first_countable_topology) first_countable_topology
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   297
proof
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   298
  fix x :: "'a \<times> 'b"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   299
  from first_countable_basisE[of "fst x"] guess A :: "'a set set" . note A = this
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   300
  from first_countable_basisE[of "snd x"] guess B :: "'b set set" . note B = this
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   301
  show "\<exists>A::('a\<times>'b) set set. countable A \<and> (\<forall>a\<in>A. x \<in> a \<and> open a) \<and> (\<forall>S. open S \<and> x \<in> S \<longrightarrow> (\<exists>a\<in>A. a \<subseteq> S))"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   302
  proof (intro exI[of _ "(\<lambda>(a, b). a \<times> b) ` (A \<times> B)"], safe)
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   303
    fix a b assume x: "a \<in> A" "b \<in> B"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   304
    with A(2, 3)[of a] B(2, 3)[of b] show "x \<in> a \<times> b" "open (a \<times> b)"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   305
      unfolding mem_Times_iff by (auto intro: open_Times)
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   306
  next
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   307
    fix S assume "open S" "x \<in> S"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   308
    from open_prod_elim[OF this] guess a' b' .
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   309
    moreover with A(4)[of a'] B(4)[of b']
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   310
    obtain a b where "a \<in> A" "a \<subseteq> a'" "b \<in> B" "b \<subseteq> b'" by auto
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   311
    ultimately show "\<exists>a\<in>(\<lambda>(a, b). a \<times> b) ` (A \<times> B). a \<subseteq> S"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   312
      by (auto intro!: bexI[of _ "a \<times> b"] bexI[of _ a] bexI[of _ b])
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   313
  qed (simp add: A B)
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   314
qed
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   315
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   316
instance metric_space \<subseteq> first_countable_topology
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   317
proof
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   318
  fix x :: 'a
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   319
  show "\<exists>A. countable A \<and> (\<forall>a\<in>A. x \<in> a \<and> open a) \<and> (\<forall>S. open S \<and> x \<in> S \<longrightarrow> (\<exists>a\<in>A. a \<subseteq> S))"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   320
  proof (intro exI, safe)
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   321
    fix S assume "open S" "x \<in> S"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   322
    then obtain r where "0 < r" "{y. dist x y < r} \<subseteq> S"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   323
      by (auto simp: open_dist dist_commute subset_eq)
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   324
    moreover from reals_Archimedean[OF `0 < r`] guess n ..
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   325
    moreover
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   326
    then have "{y. dist x y < inverse (Suc n)} \<subseteq> {y. dist x y < r}"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   327
      by (auto simp: inverse_eq_divide)
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   328
    ultimately show "\<exists>a\<in>range (\<lambda>n. {y. dist x y < inverse (Suc n)}). a \<subseteq> S"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   329
      by auto
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   330
  qed (auto intro: open_ball)
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   331
qed
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   332
50881
ae630bab13da renamed countable_basis_space to second_countable_topology
hoelzl
parents: 50526
diff changeset
   333
class second_countable_topology = topological_space +
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   334
  assumes ex_countable_basis:
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   335
    "\<exists>B::'a::topological_space set set. countable B \<and> topological_basis B"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   336
50881
ae630bab13da renamed countable_basis_space to second_countable_topology
hoelzl
parents: 50526
diff changeset
   337
sublocale second_countable_topology < countable_basis "SOME B. countable B \<and> topological_basis B"
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   338
  using someI_ex[OF ex_countable_basis] by unfold_locales safe
50094
84ddcf5364b4 allow arbitrary enumerations of basis in locale for generation of borel sets
immler
parents: 50087
diff changeset
   339
50882
a382bf90867e move prod instantiation of second_countable_topology to its definition
hoelzl
parents: 50881
diff changeset
   340
instance prod :: (second_countable_topology, second_countable_topology) second_countable_topology
a382bf90867e move prod instantiation of second_countable_topology to its definition
hoelzl
parents: 50881
diff changeset
   341
proof
a382bf90867e move prod instantiation of second_countable_topology to its definition
hoelzl
parents: 50881
diff changeset
   342
  obtain A :: "'a set set" where "countable A" "topological_basis A"
a382bf90867e move prod instantiation of second_countable_topology to its definition
hoelzl
parents: 50881
diff changeset
   343
    using ex_countable_basis by auto
a382bf90867e move prod instantiation of second_countable_topology to its definition
hoelzl
parents: 50881
diff changeset
   344
  moreover
a382bf90867e move prod instantiation of second_countable_topology to its definition
hoelzl
parents: 50881
diff changeset
   345
  obtain B :: "'b set set" where "countable B" "topological_basis B"
a382bf90867e move prod instantiation of second_countable_topology to its definition
hoelzl
parents: 50881
diff changeset
   346
    using ex_countable_basis by auto
a382bf90867e move prod instantiation of second_countable_topology to its definition
hoelzl
parents: 50881
diff changeset
   347
  ultimately show "\<exists>B::('a \<times> 'b) set set. countable B \<and> topological_basis B"
a382bf90867e move prod instantiation of second_countable_topology to its definition
hoelzl
parents: 50881
diff changeset
   348
    by (auto intro!: exI[of _ "(\<lambda>(a, b). a \<times> b) ` (A \<times> B)"] topological_basis_prod)
a382bf90867e move prod instantiation of second_countable_topology to its definition
hoelzl
parents: 50881
diff changeset
   349
qed
a382bf90867e move prod instantiation of second_countable_topology to its definition
hoelzl
parents: 50881
diff changeset
   350
50883
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   351
instance second_countable_topology \<subseteq> first_countable_topology
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   352
proof
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   353
  fix x :: 'a
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   354
  def B \<equiv> "SOME B::'a set set. countable B \<and> topological_basis B"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   355
  then have B: "countable B" "topological_basis B"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   356
    using countable_basis is_basis
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   357
    by (auto simp: countable_basis is_basis)
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   358
  then show "\<exists>A. countable A \<and> (\<forall>a\<in>A. x \<in> a \<and> open a) \<and> (\<forall>S. open S \<and> x \<in> S \<longrightarrow> (\<exists>a\<in>A. a \<subseteq> S))"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   359
    by (intro exI[of _ "{b\<in>B. x \<in> b}"])
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   360
       (fastforce simp: topological_space_class.topological_basis_def)
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   361
qed
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   362
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   363
subsection {* Polish spaces *}
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   364
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   365
text {* Textbooks define Polish spaces as completely metrizable.
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   366
  We assume the topology to be complete for a given metric. *}
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   367
50881
ae630bab13da renamed countable_basis_space to second_countable_topology
hoelzl
parents: 50526
diff changeset
   368
class polish_space = complete_space + second_countable_topology
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   369
44517
68e8eb0ce8aa minimize imports
huffman
parents: 44516
diff changeset
   370
subsection {* General notion of a topology as a value *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   371
44170
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   372
definition "istopology L \<longleftrightarrow> L {} \<and> (\<forall>S T. L S \<longrightarrow> L T \<longrightarrow> L (S \<inter> T)) \<and> (\<forall>K. Ball K L \<longrightarrow> L (\<Union> K))"
49834
b27bbb021df1 discontinued obsolete typedef (open) syntax;
wenzelm
parents: 49711
diff changeset
   373
typedef 'a topology = "{L::('a set) \<Rightarrow> bool. istopology L}"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   374
  morphisms "openin" "topology"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   375
  unfolding istopology_def by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   376
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   377
lemma istopology_open_in[intro]: "istopology(openin U)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   378
  using openin[of U] by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   379
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   380
lemma topology_inverse': "istopology U \<Longrightarrow> openin (topology U) = U"
44170
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   381
  using topology_inverse[unfolded mem_Collect_eq] .
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   382
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   383
lemma topology_inverse_iff: "istopology U \<longleftrightarrow> openin (topology U) = U"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   384
  using topology_inverse[of U] istopology_open_in[of "topology U"] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   385
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   386
lemma topology_eq: "T1 = T2 \<longleftrightarrow> (\<forall>S. openin T1 S \<longleftrightarrow> openin T2 S)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   387
proof-
49711
e5aaae7eadc9 tuned proofs;
wenzelm
parents: 48125
diff changeset
   388
  { assume "T1=T2"
e5aaae7eadc9 tuned proofs;
wenzelm
parents: 48125
diff changeset
   389
    hence "\<forall>S. openin T1 S \<longleftrightarrow> openin T2 S" by simp }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   390
  moreover
49711
e5aaae7eadc9 tuned proofs;
wenzelm
parents: 48125
diff changeset
   391
  { assume H: "\<forall>S. openin T1 S \<longleftrightarrow> openin T2 S"
44170
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   392
    hence "openin T1 = openin T2" by (simp add: fun_eq_iff)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   393
    hence "topology (openin T1) = topology (openin T2)" by simp
49711
e5aaae7eadc9 tuned proofs;
wenzelm
parents: 48125
diff changeset
   394
    hence "T1 = T2" unfolding openin_inverse .
e5aaae7eadc9 tuned proofs;
wenzelm
parents: 48125
diff changeset
   395
  }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   396
  ultimately show ?thesis by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   397
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   398
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   399
text{* Infer the "universe" from union of all sets in the topology. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   400
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   401
definition "topspace T =  \<Union>{S. openin T S}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   402
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   403
subsubsection {* Main properties of open sets *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   404
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   405
lemma openin_clauses:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   406
  fixes U :: "'a topology"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   407
  shows "openin U {}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   408
  "\<And>S T. openin U S \<Longrightarrow> openin U T \<Longrightarrow> openin U (S\<inter>T)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   409
  "\<And>K. (\<forall>S \<in> K. openin U S) \<Longrightarrow> openin U (\<Union>K)"
44170
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   410
  using openin[of U] unfolding istopology_def mem_Collect_eq
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   411
  by fast+
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   412
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   413
lemma openin_subset[intro]: "openin U S \<Longrightarrow> S \<subseteq> topspace U"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   414
  unfolding topspace_def by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   415
lemma openin_empty[simp]: "openin U {}" by (simp add: openin_clauses)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   416
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   417
lemma openin_Int[intro]: "openin U S \<Longrightarrow> openin U T \<Longrightarrow> openin U (S \<inter> T)"
36362
06475a1547cb fix lots of looping simp calls and other warnings
huffman
parents: 36360
diff changeset
   418
  using openin_clauses by simp
06475a1547cb fix lots of looping simp calls and other warnings
huffman
parents: 36360
diff changeset
   419
06475a1547cb fix lots of looping simp calls and other warnings
huffman
parents: 36360
diff changeset
   420
lemma openin_Union[intro]: "(\<forall>S \<in>K. openin U S) \<Longrightarrow> openin U (\<Union> K)"
06475a1547cb fix lots of looping simp calls and other warnings
huffman
parents: 36360
diff changeset
   421
  using openin_clauses by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   422
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   423
lemma openin_Un[intro]: "openin U S \<Longrightarrow> openin U T \<Longrightarrow> openin U (S \<union> T)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   424
  using openin_Union[of "{S,T}" U] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   425
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   426
lemma openin_topspace[intro, simp]: "openin U (topspace U)" by (simp add: openin_Union topspace_def)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   427
49711
e5aaae7eadc9 tuned proofs;
wenzelm
parents: 48125
diff changeset
   428
lemma openin_subopen: "openin U S \<longleftrightarrow> (\<forall>x \<in> S. \<exists>T. openin U T \<and> x \<in> T \<and> T \<subseteq> S)"
e5aaae7eadc9 tuned proofs;
wenzelm
parents: 48125
diff changeset
   429
  (is "?lhs \<longleftrightarrow> ?rhs")
36584
1535841fc2e9 prove lemma openin_subopen without using choice
huffman
parents: 36442
diff changeset
   430
proof
49711
e5aaae7eadc9 tuned proofs;
wenzelm
parents: 48125
diff changeset
   431
  assume ?lhs
e5aaae7eadc9 tuned proofs;
wenzelm
parents: 48125
diff changeset
   432
  then show ?rhs by auto
36584
1535841fc2e9 prove lemma openin_subopen without using choice
huffman
parents: 36442
diff changeset
   433
next
1535841fc2e9 prove lemma openin_subopen without using choice
huffman
parents: 36442
diff changeset
   434
  assume H: ?rhs
1535841fc2e9 prove lemma openin_subopen without using choice
huffman
parents: 36442
diff changeset
   435
  let ?t = "\<Union>{T. openin U T \<and> T \<subseteq> S}"
1535841fc2e9 prove lemma openin_subopen without using choice
huffman
parents: 36442
diff changeset
   436
  have "openin U ?t" by (simp add: openin_Union)
1535841fc2e9 prove lemma openin_subopen without using choice
huffman
parents: 36442
diff changeset
   437
  also have "?t = S" using H by auto
1535841fc2e9 prove lemma openin_subopen without using choice
huffman
parents: 36442
diff changeset
   438
  finally show "openin U S" .
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   439
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   440
49711
e5aaae7eadc9 tuned proofs;
wenzelm
parents: 48125
diff changeset
   441
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   442
subsubsection {* Closed sets *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   443
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   444
definition "closedin U S \<longleftrightarrow> S \<subseteq> topspace U \<and> openin U (topspace U - S)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   445
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   446
lemma closedin_subset: "closedin U S \<Longrightarrow> S \<subseteq> topspace U" by (metis closedin_def)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   447
lemma closedin_empty[simp]: "closedin U {}" by (simp add: closedin_def)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   448
lemma closedin_topspace[intro,simp]:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   449
  "closedin U (topspace U)" by (simp add: closedin_def)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   450
lemma closedin_Un[intro]: "closedin U S \<Longrightarrow> closedin U T \<Longrightarrow> closedin U (S \<union> T)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   451
  by (auto simp add: Diff_Un closedin_def)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   452
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   453
lemma Diff_Inter[intro]: "A - \<Inter>S = \<Union> {A - s|s. s\<in>S}" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   454
lemma closedin_Inter[intro]: assumes Ke: "K \<noteq> {}" and Kc: "\<forall>S \<in>K. closedin U S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   455
  shows "closedin U (\<Inter> K)"  using Ke Kc unfolding closedin_def Diff_Inter by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   456
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   457
lemma closedin_Int[intro]: "closedin U S \<Longrightarrow> closedin U T \<Longrightarrow> closedin U (S \<inter> T)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   458
  using closedin_Inter[of "{S,T}" U] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   459
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   460
lemma Diff_Diff_Int: "A - (A - B) = A \<inter> B" by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   461
lemma openin_closedin_eq: "openin U S \<longleftrightarrow> S \<subseteq> topspace U \<and> closedin U (topspace U - S)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   462
  apply (auto simp add: closedin_def Diff_Diff_Int inf_absorb2)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   463
  apply (metis openin_subset subset_eq)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   464
  done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   465
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   466
lemma openin_closedin:  "S \<subseteq> topspace U \<Longrightarrow> (openin U S \<longleftrightarrow> closedin U (topspace U - S))"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   467
  by (simp add: openin_closedin_eq)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   468
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   469
lemma openin_diff[intro]: assumes oS: "openin U S" and cT: "closedin U T" shows "openin U (S - T)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   470
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   471
  have "S - T = S \<inter> (topspace U - T)" using openin_subset[of U S]  oS cT
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   472
    by (auto simp add: topspace_def openin_subset)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   473
  then show ?thesis using oS cT by (auto simp add: closedin_def)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   474
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   475
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   476
lemma closedin_diff[intro]: assumes oS: "closedin U S" and cT: "openin U T" shows "closedin U (S - T)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   477
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   478
  have "S - T = S \<inter> (topspace U - T)" using closedin_subset[of U S]  oS cT
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   479
    by (auto simp add: topspace_def )
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   480
  then show ?thesis using oS cT by (auto simp add: openin_closedin_eq)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   481
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   482
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   483
subsubsection {* Subspace topology *}
44170
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   484
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   485
definition "subtopology U V = topology (\<lambda>T. \<exists>S. T = S \<inter> V \<and> openin U S)"
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   486
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   487
lemma istopology_subtopology: "istopology (\<lambda>T. \<exists>S. T = S \<inter> V \<and> openin U S)"
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   488
  (is "istopology ?L")
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   489
proof-
44170
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   490
  have "?L {}" by blast
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   491
  {fix A B assume A: "?L A" and B: "?L B"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   492
    from A B obtain Sa and Sb where Sa: "openin U Sa" "A = Sa \<inter> V" and Sb: "openin U Sb" "B = Sb \<inter> V" by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   493
    have "A\<inter>B = (Sa \<inter> Sb) \<inter> V" "openin U (Sa \<inter> Sb)"  using Sa Sb by blast+
44170
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   494
    then have "?L (A \<inter> B)" by blast}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   495
  moreover
44170
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   496
  {fix K assume K: "K \<subseteq> Collect ?L"
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   497
    have th0: "Collect ?L = (\<lambda>S. S \<inter> V) ` Collect (openin U)"
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
   498
      apply (rule set_eqI)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   499
      apply (simp add: Ball_def image_iff)
44170
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   500
      by metis
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   501
    from K[unfolded th0 subset_image_iff]
44170
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   502
    obtain Sk where Sk: "Sk \<subseteq> Collect (openin U)" "K = (\<lambda>S. S \<inter> V) ` Sk" by blast
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   503
    have "\<Union>K = (\<Union>Sk) \<inter> V" using Sk by auto
44170
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   504
    moreover have "openin U (\<Union> Sk)" using Sk by (auto simp add: subset_eq)
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   505
    ultimately have "?L (\<Union>K)" by blast}
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   506
  ultimately show ?thesis
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   507
    unfolding subset_eq mem_Collect_eq istopology_def by blast
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   508
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   509
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   510
lemma openin_subtopology:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   511
  "openin (subtopology U V) S \<longleftrightarrow> (\<exists> T. (openin U T) \<and> (S = T \<inter> V))"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   512
  unfolding subtopology_def topology_inverse'[OF istopology_subtopology]
44170
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   513
  by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   514
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   515
lemma topspace_subtopology: "topspace(subtopology U V) = topspace U \<inter> V"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   516
  by (auto simp add: topspace_def openin_subtopology)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   517
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   518
lemma closedin_subtopology:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   519
  "closedin (subtopology U V) S \<longleftrightarrow> (\<exists>T. closedin U T \<and> S = T \<inter> V)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   520
  unfolding closedin_def topspace_subtopology
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   521
  apply (simp add: openin_subtopology)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   522
  apply (rule iffI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   523
  apply clarify
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   524
  apply (rule_tac x="topspace U - T" in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   525
  by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   526
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   527
lemma openin_subtopology_refl: "openin (subtopology U V) V \<longleftrightarrow> V \<subseteq> topspace U"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   528
  unfolding openin_subtopology
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   529
  apply (rule iffI, clarify)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   530
  apply (frule openin_subset[of U])  apply blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   531
  apply (rule exI[where x="topspace U"])
49711
e5aaae7eadc9 tuned proofs;
wenzelm
parents: 48125
diff changeset
   532
  apply auto
e5aaae7eadc9 tuned proofs;
wenzelm
parents: 48125
diff changeset
   533
  done
e5aaae7eadc9 tuned proofs;
wenzelm
parents: 48125
diff changeset
   534
e5aaae7eadc9 tuned proofs;
wenzelm
parents: 48125
diff changeset
   535
lemma subtopology_superset:
e5aaae7eadc9 tuned proofs;
wenzelm
parents: 48125
diff changeset
   536
  assumes UV: "topspace U \<subseteq> V"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   537
  shows "subtopology U V = U"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   538
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   539
  {fix S
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   540
    {fix T assume T: "openin U T" "S = T \<inter> V"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   541
      from T openin_subset[OF T(1)] UV have eq: "S = T" by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   542
      have "openin U S" unfolding eq using T by blast}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   543
    moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   544
    {assume S: "openin U S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   545
      hence "\<exists>T. openin U T \<and> S = T \<inter> V"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   546
        using openin_subset[OF S] UV by auto}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   547
    ultimately have "(\<exists>T. openin U T \<and> S = T \<inter> V) \<longleftrightarrow> openin U S" by blast}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   548
  then show ?thesis unfolding topology_eq openin_subtopology by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   549
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   550
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   551
lemma subtopology_topspace[simp]: "subtopology U (topspace U) = U"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   552
  by (simp add: subtopology_superset)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   553
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   554
lemma subtopology_UNIV[simp]: "subtopology U UNIV = U"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   555
  by (simp add: subtopology_superset)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   556
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   557
subsubsection {* The standard Euclidean topology *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   558
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   559
definition
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   560
  euclidean :: "'a::topological_space topology" where
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   561
  "euclidean = topology open"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   562
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   563
lemma open_openin: "open S \<longleftrightarrow> openin euclidean S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   564
  unfolding euclidean_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   565
  apply (rule cong[where x=S and y=S])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   566
  apply (rule topology_inverse[symmetric])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   567
  apply (auto simp add: istopology_def)
44170
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   568
  done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   569
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   570
lemma topspace_euclidean: "topspace euclidean = UNIV"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   571
  apply (simp add: topspace_def)
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
   572
  apply (rule set_eqI)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   573
  by (auto simp add: open_openin[symmetric])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   574
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   575
lemma topspace_euclidean_subtopology[simp]: "topspace (subtopology euclidean S) = S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   576
  by (simp add: topspace_euclidean topspace_subtopology)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   577
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   578
lemma closed_closedin: "closed S \<longleftrightarrow> closedin euclidean S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   579
  by (simp add: closed_def closedin_def topspace_euclidean open_openin Compl_eq_Diff_UNIV)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   580
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   581
lemma open_subopen: "open S \<longleftrightarrow> (\<forall>x\<in>S. \<exists>T. open T \<and> x \<in> T \<and> T \<subseteq> S)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   582
  by (simp add: open_openin openin_subopen[symmetric])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   583
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   584
text {* Basic "localization" results are handy for connectedness. *}
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   585
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   586
lemma openin_open: "openin (subtopology euclidean U) S \<longleftrightarrow> (\<exists>T. open T \<and> (S = U \<inter> T))"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   587
  by (auto simp add: openin_subtopology open_openin[symmetric])
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   588
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   589
lemma openin_open_Int[intro]: "open S \<Longrightarrow> openin (subtopology euclidean U) (U \<inter> S)"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   590
  by (auto simp add: openin_open)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   591
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   592
lemma open_openin_trans[trans]:
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   593
 "open S \<Longrightarrow> open T \<Longrightarrow> T \<subseteq> S \<Longrightarrow> openin (subtopology euclidean S) T"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   594
  by (metis Int_absorb1  openin_open_Int)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   595
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   596
lemma open_subset:  "S \<subseteq> T \<Longrightarrow> open S \<Longrightarrow> openin (subtopology euclidean T) S"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   597
  by (auto simp add: openin_open)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   598
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   599
lemma closedin_closed: "closedin (subtopology euclidean U) S \<longleftrightarrow> (\<exists>T. closed T \<and> S = U \<inter> T)"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   600
  by (simp add: closedin_subtopology closed_closedin Int_ac)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   601
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   602
lemma closedin_closed_Int: "closed S ==> closedin (subtopology euclidean U) (U \<inter> S)"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   603
  by (metis closedin_closed)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   604
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   605
lemma closed_closedin_trans: "closed S \<Longrightarrow> closed T \<Longrightarrow> T \<subseteq> S \<Longrightarrow> closedin (subtopology euclidean S) T"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   606
  apply (subgoal_tac "S \<inter> T = T" )
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   607
  apply auto
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   608
  apply (frule closedin_closed_Int[of T S])
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   609
  by simp
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   610
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   611
lemma closed_subset: "S \<subseteq> T \<Longrightarrow> closed S \<Longrightarrow> closedin (subtopology euclidean T) S"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   612
  by (auto simp add: closedin_closed)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   613
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   614
lemma openin_euclidean_subtopology_iff:
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   615
  fixes S U :: "'a::metric_space set"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   616
  shows "openin (subtopology euclidean U) S
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   617
  \<longleftrightarrow> S \<subseteq> U \<and> (\<forall>x\<in>S. \<exists>e>0. \<forall>x'\<in>U. dist x' x < e \<longrightarrow> x'\<in> S)" (is "?lhs \<longleftrightarrow> ?rhs")
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   618
proof
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   619
  assume ?lhs thus ?rhs unfolding openin_open open_dist by blast
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   620
next
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   621
  def T \<equiv> "{x. \<exists>a\<in>S. \<exists>d>0. (\<forall>y\<in>U. dist y a < d \<longrightarrow> y \<in> S) \<and> dist x a < d}"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   622
  have 1: "\<forall>x\<in>T. \<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> T"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   623
    unfolding T_def
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   624
    apply clarsimp
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   625
    apply (rule_tac x="d - dist x a" in exI)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   626
    apply (clarsimp simp add: less_diff_eq)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   627
    apply (erule rev_bexI)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   628
    apply (rule_tac x=d in exI, clarify)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   629
    apply (erule le_less_trans [OF dist_triangle])
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   630
    done
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   631
  assume ?rhs hence 2: "S = U \<inter> T"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   632
    unfolding T_def
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   633
    apply auto
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   634
    apply (drule (1) bspec, erule rev_bexI)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   635
    apply auto
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   636
    done
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   637
  from 1 2 show ?lhs
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   638
    unfolding openin_open open_dist by fast
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   639
qed
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   640
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   641
text {* These "transitivity" results are handy too *}
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   642
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   643
lemma openin_trans[trans]: "openin (subtopology euclidean T) S \<Longrightarrow> openin (subtopology euclidean U) T
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   644
  \<Longrightarrow> openin (subtopology euclidean U) S"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   645
  unfolding open_openin openin_open by blast
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   646
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   647
lemma openin_open_trans: "openin (subtopology euclidean T) S \<Longrightarrow> open T \<Longrightarrow> open S"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   648
  by (auto simp add: openin_open intro: openin_trans)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   649
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   650
lemma closedin_trans[trans]:
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   651
 "closedin (subtopology euclidean T) S \<Longrightarrow>
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   652
           closedin (subtopology euclidean U) T
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   653
           ==> closedin (subtopology euclidean U) S"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   654
  by (auto simp add: closedin_closed closed_closedin closed_Inter Int_assoc)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   655
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   656
lemma closedin_closed_trans: "closedin (subtopology euclidean T) S \<Longrightarrow> closed T \<Longrightarrow> closed S"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   657
  by (auto simp add: closedin_closed intro: closedin_trans)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   658
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   659
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   660
subsection {* Open and closed balls *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   661
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   662
definition
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   663
  ball :: "'a::metric_space \<Rightarrow> real \<Rightarrow> 'a set" where
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   664
  "ball x e = {y. dist x y < e}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   665
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   666
definition
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   667
  cball :: "'a::metric_space \<Rightarrow> real \<Rightarrow> 'a set" where
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   668
  "cball x e = {y. dist x y \<le> e}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   669
45776
714100f5fda4 remove mem_(c)ball_0 and centre_in_(c)ball from simpset, as rules mem_(c)ball always match instead
huffman
parents: 45548
diff changeset
   670
lemma mem_ball [simp]: "y \<in> ball x e \<longleftrightarrow> dist x y < e"
714100f5fda4 remove mem_(c)ball_0 and centre_in_(c)ball from simpset, as rules mem_(c)ball always match instead
huffman
parents: 45548
diff changeset
   671
  by (simp add: ball_def)
714100f5fda4 remove mem_(c)ball_0 and centre_in_(c)ball from simpset, as rules mem_(c)ball always match instead
huffman
parents: 45548
diff changeset
   672
714100f5fda4 remove mem_(c)ball_0 and centre_in_(c)ball from simpset, as rules mem_(c)ball always match instead
huffman
parents: 45548
diff changeset
   673
lemma mem_cball [simp]: "y \<in> cball x e \<longleftrightarrow> dist x y \<le> e"
714100f5fda4 remove mem_(c)ball_0 and centre_in_(c)ball from simpset, as rules mem_(c)ball always match instead
huffman
parents: 45548
diff changeset
   674
  by (simp add: cball_def)
714100f5fda4 remove mem_(c)ball_0 and centre_in_(c)ball from simpset, as rules mem_(c)ball always match instead
huffman
parents: 45548
diff changeset
   675
714100f5fda4 remove mem_(c)ball_0 and centre_in_(c)ball from simpset, as rules mem_(c)ball always match instead
huffman
parents: 45548
diff changeset
   676
lemma mem_ball_0:
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   677
  fixes x :: "'a::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   678
  shows "x \<in> ball 0 e \<longleftrightarrow> norm x < e"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   679
  by (simp add: dist_norm)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   680
45776
714100f5fda4 remove mem_(c)ball_0 and centre_in_(c)ball from simpset, as rules mem_(c)ball always match instead
huffman
parents: 45548
diff changeset
   681
lemma mem_cball_0:
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   682
  fixes x :: "'a::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   683
  shows "x \<in> cball 0 e \<longleftrightarrow> norm x \<le> e"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   684
  by (simp add: dist_norm)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   685
45776
714100f5fda4 remove mem_(c)ball_0 and centre_in_(c)ball from simpset, as rules mem_(c)ball always match instead
huffman
parents: 45548
diff changeset
   686
lemma centre_in_ball: "x \<in> ball x e \<longleftrightarrow> 0 < e"
714100f5fda4 remove mem_(c)ball_0 and centre_in_(c)ball from simpset, as rules mem_(c)ball always match instead
huffman
parents: 45548
diff changeset
   687
  by simp
714100f5fda4 remove mem_(c)ball_0 and centre_in_(c)ball from simpset, as rules mem_(c)ball always match instead
huffman
parents: 45548
diff changeset
   688
714100f5fda4 remove mem_(c)ball_0 and centre_in_(c)ball from simpset, as rules mem_(c)ball always match instead
huffman
parents: 45548
diff changeset
   689
lemma centre_in_cball: "x \<in> cball x e \<longleftrightarrow> 0 \<le> e"
714100f5fda4 remove mem_(c)ball_0 and centre_in_(c)ball from simpset, as rules mem_(c)ball always match instead
huffman
parents: 45548
diff changeset
   690
  by simp
714100f5fda4 remove mem_(c)ball_0 and centre_in_(c)ball from simpset, as rules mem_(c)ball always match instead
huffman
parents: 45548
diff changeset
   691
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   692
lemma ball_subset_cball[simp,intro]: "ball x e \<subseteq> cball x e" by (simp add: subset_eq)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   693
lemma subset_ball[intro]: "d <= e ==> ball x d \<subseteq> ball x e" by (simp add: subset_eq)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   694
lemma subset_cball[intro]: "d <= e ==> cball x d \<subseteq> cball x e" by (simp add: subset_eq)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   695
lemma ball_max_Un: "ball a (max r s) = ball a r \<union> ball a s"
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
   696
  by (simp add: set_eq_iff) arith
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   697
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   698
lemma ball_min_Int: "ball a (min r s) = ball a r \<inter> ball a s"
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
   699
  by (simp add: set_eq_iff)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   700
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   701
lemma diff_less_iff: "(a::real) - b > 0 \<longleftrightarrow> a > b"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   702
  "(a::real) - b < 0 \<longleftrightarrow> a < b"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   703
  "a - b < c \<longleftrightarrow> a < c +b" "a - b > c \<longleftrightarrow> a > c +b" by arith+
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   704
lemma diff_le_iff: "(a::real) - b \<ge> 0 \<longleftrightarrow> a \<ge> b" "(a::real) - b \<le> 0 \<longleftrightarrow> a \<le> b"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   705
  "a - b \<le> c \<longleftrightarrow> a \<le> c +b" "a - b \<ge> c \<longleftrightarrow> a \<ge> c +b"  by arith+
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   706
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   707
lemma open_ball[intro, simp]: "open (ball x e)"
44170
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   708
  unfolding open_dist ball_def mem_Collect_eq Ball_def
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   709
  unfolding dist_commute
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   710
  apply clarify
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   711
  apply (rule_tac x="e - dist xa x" in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   712
  using dist_triangle_alt[where z=x]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   713
  apply (clarsimp simp add: diff_less_iff)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   714
  apply atomize
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   715
  apply (erule_tac x="y" in allE)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   716
  apply (erule_tac x="xa" in allE)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   717
  by arith
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   718
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   719
lemma open_contains_ball: "open S \<longleftrightarrow> (\<forall>x\<in>S. \<exists>e>0. ball x e \<subseteq> S)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   720
  unfolding open_dist subset_eq mem_ball Ball_def dist_commute ..
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   721
33714
eb2574ac4173 Added new lemmas to Euclidean Space by Robert Himmelmann
hoelzl
parents: 33324
diff changeset
   722
lemma openE[elim?]:
eb2574ac4173 Added new lemmas to Euclidean Space by Robert Himmelmann
hoelzl
parents: 33324
diff changeset
   723
  assumes "open S" "x\<in>S" 
eb2574ac4173 Added new lemmas to Euclidean Space by Robert Himmelmann
hoelzl
parents: 33324
diff changeset
   724
  obtains e where "e>0" "ball x e \<subseteq> S"
eb2574ac4173 Added new lemmas to Euclidean Space by Robert Himmelmann
hoelzl
parents: 33324
diff changeset
   725
  using assms unfolding open_contains_ball by auto
eb2574ac4173 Added new lemmas to Euclidean Space by Robert Himmelmann
hoelzl
parents: 33324
diff changeset
   726
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   727
lemma open_contains_ball_eq: "open S \<Longrightarrow> \<forall>x. x\<in>S \<longleftrightarrow> (\<exists>e>0. ball x e \<subseteq> S)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   728
  by (metis open_contains_ball subset_eq centre_in_ball)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   729
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   730
lemma ball_eq_empty[simp]: "ball x e = {} \<longleftrightarrow> e \<le> 0"
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
   731
  unfolding mem_ball set_eq_iff
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   732
  apply (simp add: not_less)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   733
  by (metis zero_le_dist order_trans dist_self)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   734
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   735
lemma ball_empty[intro]: "e \<le> 0 ==> ball x e = {}" by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   736
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
   737
lemma euclidean_dist_l2:
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
   738
  fixes x y :: "'a :: euclidean_space"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
   739
  shows "dist x y = setL2 (\<lambda>i. dist (x \<bullet> i) (y \<bullet> i)) Basis"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
   740
  unfolding dist_norm norm_eq_sqrt_inner setL2_def
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
   741
  by (subst euclidean_inner) (simp add: power2_eq_square inner_diff_left)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
   742
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
   743
definition "box a b = {x. \<forall>i\<in>Basis. a \<bullet> i < x \<bullet> i \<and> x \<bullet> i < b \<bullet> i}"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
   744
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   745
lemma rational_boxes:
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
   746
  fixes x :: "'a\<Colon>euclidean_space"
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   747
  assumes "0 < e"
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
   748
  shows "\<exists>a b. (\<forall>i\<in>Basis. a \<bullet> i \<in> \<rat> \<and> b \<bullet> i \<in> \<rat> ) \<and> x \<in> box a b \<and> box a b \<subseteq> ball x e"
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   749
proof -
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   750
  def e' \<equiv> "e / (2 * sqrt (real (DIM ('a))))"
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
   751
  then have e: "0 < e'" using assms by (auto intro!: divide_pos_pos simp: DIM_positive)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
   752
  have "\<forall>i. \<exists>y. y \<in> \<rat> \<and> y < x \<bullet> i \<and> x \<bullet> i - y < e'" (is "\<forall>i. ?th i")
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   753
  proof
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
   754
    fix i from Rats_dense_in_real[of "x \<bullet> i - e'" "x \<bullet> i"] e show "?th i" by auto
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   755
  qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   756
  from choice[OF this] guess a .. note a = this
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
   757
  have "\<forall>i. \<exists>y. y \<in> \<rat> \<and> x \<bullet> i < y \<and> y - x \<bullet> i < e'" (is "\<forall>i. ?th i")
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   758
  proof
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
   759
    fix i from Rats_dense_in_real[of "x \<bullet> i" "x \<bullet> i + e'"] e show "?th i" by auto
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   760
  qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   761
  from choice[OF this] guess b .. note b = this
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
   762
  let ?a = "\<Sum>i\<in>Basis. a i *\<^sub>R i" and ?b = "\<Sum>i\<in>Basis. b i *\<^sub>R i"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
   763
  show ?thesis
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
   764
  proof (rule exI[of _ ?a], rule exI[of _ ?b], safe)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
   765
    fix y :: 'a assume *: "y \<in> box ?a ?b"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
   766
    have "dist x y = sqrt (\<Sum>i\<in>Basis. (dist (x \<bullet> i) (y \<bullet> i))\<twosuperior>)"
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   767
      unfolding setL2_def[symmetric] by (rule euclidean_dist_l2)
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
   768
    also have "\<dots> < sqrt (\<Sum>(i::'a)\<in>Basis. e^2 / real (DIM('a)))"
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   769
    proof (rule real_sqrt_less_mono, rule setsum_strict_mono)
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
   770
      fix i :: "'a" assume i: "i \<in> Basis"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
   771
      have "a i < y\<bullet>i \<and> y\<bullet>i < b i" using * i by (auto simp: box_def)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
   772
      moreover have "a i < x\<bullet>i" "x\<bullet>i - a i < e'" using a by auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
   773
      moreover have "x\<bullet>i < b i" "b i - x\<bullet>i < e'" using b by auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
   774
      ultimately have "\<bar>x\<bullet>i - y\<bullet>i\<bar> < 2 * e'" by auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
   775
      then have "dist (x \<bullet> i) (y \<bullet> i) < e/sqrt (real (DIM('a)))"
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   776
        unfolding e'_def by (auto simp: dist_real_def)
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
   777
      then have "(dist (x \<bullet> i) (y \<bullet> i))\<twosuperior> < (e/sqrt (real (DIM('a))))\<twosuperior>"
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   778
        by (rule power_strict_mono) auto
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
   779
      then show "(dist (x \<bullet> i) (y \<bullet> i))\<twosuperior> < e\<twosuperior> / real DIM('a)"
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   780
        by (simp add: power_divide)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   781
    qed auto
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
   782
    also have "\<dots> = e" using `0 < e` by (simp add: real_eq_of_nat)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
   783
    finally show "y \<in> ball x e" by (auto simp: ball_def)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
   784
  qed (insert a b, auto simp: box_def)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
   785
qed
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
   786
 
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
   787
lemma open_UNION_box:
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
   788
  fixes M :: "'a\<Colon>euclidean_space set"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
   789
  assumes "open M" 
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
   790
  defines "a' \<equiv> \<lambda>f :: 'a \<Rightarrow> real \<times> real. (\<Sum>(i::'a)\<in>Basis. fst (f i) *\<^sub>R i)"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
   791
  defines "b' \<equiv> \<lambda>f :: 'a \<Rightarrow> real \<times> real. (\<Sum>(i::'a)\<in>Basis. snd (f i) *\<^sub>R i)"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
   792
  defines "I \<equiv> {f\<in>Basis \<rightarrow>\<^isub>E \<rat> \<times> \<rat>. box (a' f) (b' f) \<subseteq> M}"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
   793
  shows "M = (\<Union>f\<in>I. box (a' f) (b' f))"
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   794
proof safe
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   795
  fix x assume "x \<in> M"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   796
  obtain e where e: "e > 0" "ball x e \<subseteq> M"
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
   797
    using openE[OF `open M` `x \<in> M`] by auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
   798
  moreover then obtain a b where ab: "x \<in> box a b"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
   799
    "\<forall>i \<in> Basis. a \<bullet> i \<in> \<rat>" "\<forall>i\<in>Basis. b \<bullet> i \<in> \<rat>" "box a b \<subseteq> ball x e"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
   800
    using rational_boxes[OF e(1)] by metis
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
   801
  ultimately show "x \<in> (\<Union>f\<in>I. box (a' f) (b' f))"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
   802
     by (intro UN_I[of "\<lambda>i\<in>Basis. (a \<bullet> i, b \<bullet> i)"])
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
   803
        (auto simp: euclidean_representation I_def a'_def b'_def)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
   804
qed (auto simp: I_def)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   805
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   806
subsection{* Connectedness *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   807
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   808
definition "connected S \<longleftrightarrow>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   809
  ~(\<exists>e1 e2. open e1 \<and> open e2 \<and> S \<subseteq> (e1 \<union> e2) \<and> (e1 \<inter> e2 \<inter> S = {})
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   810
  \<and> ~(e1 \<inter> S = {}) \<and> ~(e2 \<inter> S = {}))"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   811
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   812
lemma connected_local:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   813
 "connected S \<longleftrightarrow> ~(\<exists>e1 e2.
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   814
                 openin (subtopology euclidean S) e1 \<and>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   815
                 openin (subtopology euclidean S) e2 \<and>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   816
                 S \<subseteq> e1 \<union> e2 \<and>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   817
                 e1 \<inter> e2 = {} \<and>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   818
                 ~(e1 = {}) \<and>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   819
                 ~(e2 = {}))"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   820
unfolding connected_def openin_open by (safe, blast+)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   821
34105
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
   822
lemma exists_diff:
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
   823
  fixes P :: "'a set \<Rightarrow> bool"
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
   824
  shows "(\<exists>S. P(- S)) \<longleftrightarrow> (\<exists>S. P S)" (is "?lhs \<longleftrightarrow> ?rhs")
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   825
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   826
  {assume "?lhs" hence ?rhs by blast }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   827
  moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   828
  {fix S assume H: "P S"
34105
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
   829
    have "S = - (- S)" by auto
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
   830
    with H have "P (- (- S))" by metis }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   831
  ultimately show ?thesis by metis
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   832
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   833
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   834
lemma connected_clopen: "connected S \<longleftrightarrow>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   835
        (\<forall>T. openin (subtopology euclidean S) T \<and>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   836
            closedin (subtopology euclidean S) T \<longrightarrow> T = {} \<or> T = S)" (is "?lhs \<longleftrightarrow> ?rhs")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   837
proof-
34105
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
   838
  have " \<not> connected S \<longleftrightarrow> (\<exists>e1 e2. open e1 \<and> open (- e2) \<and> S \<subseteq> e1 \<union> (- e2) \<and> e1 \<inter> (- e2) \<inter> S = {} \<and> e1 \<inter> S \<noteq> {} \<and> (- e2) \<inter> S \<noteq> {})"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   839
    unfolding connected_def openin_open closedin_closed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   840
    apply (subst exists_diff) by blast
34105
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
   841
  hence th0: "connected S \<longleftrightarrow> \<not> (\<exists>e2 e1. closed e2 \<and> open e1 \<and> S \<subseteq> e1 \<union> (- e2) \<and> e1 \<inter> (- e2) \<inter> S = {} \<and> e1 \<inter> S \<noteq> {} \<and> (- e2) \<inter> S \<noteq> {})"
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
   842
    (is " _ \<longleftrightarrow> \<not> (\<exists>e2 e1. ?P e2 e1)") apply (simp add: closed_def) by metis
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   843
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   844
  have th1: "?rhs \<longleftrightarrow> \<not> (\<exists>t' t. closed t'\<and>t = S\<inter>t' \<and> t\<noteq>{} \<and> t\<noteq>S \<and> (\<exists>t'. open t' \<and> t = S \<inter> t'))"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   845
    (is "_ \<longleftrightarrow> \<not> (\<exists>t' t. ?Q t' t)")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   846
    unfolding connected_def openin_open closedin_closed by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   847
  {fix e2
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   848
    {fix e1 have "?P e2 e1 \<longleftrightarrow> (\<exists>t.  closed e2 \<and> t = S\<inter>e2 \<and> open e1 \<and> t = S\<inter>e1 \<and> t\<noteq>{} \<and> t\<noteq>S)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   849
        by auto}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   850
    then have "(\<exists>e1. ?P e2 e1) \<longleftrightarrow> (\<exists>t. ?Q e2 t)" by metis}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   851
  then have "\<forall>e2. (\<exists>e1. ?P e2 e1) \<longleftrightarrow> (\<exists>t. ?Q e2 t)" by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   852
  then show ?thesis unfolding th0 th1 by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   853
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   854
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   855
lemma connected_empty[simp, intro]: "connected {}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   856
  by (simp add: connected_def)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   857
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   858
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   859
subsection{* Limit points *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   860
44207
ea99698c2070 locale-ize some definitions, so perfect_space and heine_borel can inherit from the proper superclasses
huffman
parents: 44170
diff changeset
   861
definition (in topological_space)
ea99698c2070 locale-ize some definitions, so perfect_space and heine_borel can inherit from the proper superclasses
huffman
parents: 44170
diff changeset
   862
  islimpt:: "'a \<Rightarrow> 'a set \<Rightarrow> bool" (infixr "islimpt" 60) where
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   863
  "x islimpt S \<longleftrightarrow> (\<forall>T. x\<in>T \<longrightarrow> open T \<longrightarrow> (\<exists>y\<in>S. y\<in>T \<and> y\<noteq>x))"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   864
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   865
lemma islimptI:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   866
  assumes "\<And>T. x \<in> T \<Longrightarrow> open T \<Longrightarrow> \<exists>y\<in>S. y \<in> T \<and> y \<noteq> x"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   867
  shows "x islimpt S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   868
  using assms unfolding islimpt_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   869
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   870
lemma islimptE:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   871
  assumes "x islimpt S" and "x \<in> T" and "open T"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   872
  obtains y where "y \<in> S" and "y \<in> T" and "y \<noteq> x"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   873
  using assms unfolding islimpt_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   874
44584
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
   875
lemma islimpt_iff_eventually: "x islimpt S \<longleftrightarrow> \<not> eventually (\<lambda>y. y \<notin> S) (at x)"
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
   876
  unfolding islimpt_def eventually_at_topological by auto
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
   877
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
   878
lemma islimpt_subset: "\<lbrakk>x islimpt S; S \<subseteq> T\<rbrakk> \<Longrightarrow> x islimpt T"
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
   879
  unfolding islimpt_def by fast
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   880
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   881
lemma islimpt_approachable:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   882
  fixes x :: "'a::metric_space"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   883
  shows "x islimpt S \<longleftrightarrow> (\<forall>e>0. \<exists>x'\<in>S. x' \<noteq> x \<and> dist x' x < e)"
44584
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
   884
  unfolding islimpt_iff_eventually eventually_at by fast
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   885
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   886
lemma islimpt_approachable_le:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   887
  fixes x :: "'a::metric_space"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   888
  shows "x islimpt S \<longleftrightarrow> (\<forall>e>0. \<exists>x'\<in> S. x' \<noteq> x \<and> dist x' x <= e)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   889
  unfolding islimpt_approachable
44584
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
   890
  using approachable_lt_le [where f="\<lambda>y. dist y x" and P="\<lambda>y. y \<notin> S \<or> y = x",
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
   891
    THEN arg_cong [where f=Not]]
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
   892
  by (simp add: Bex_def conj_commute conj_left_commute)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   893
44571
bd91b77c4cd6 move class perfect_space into RealVector.thy;
huffman
parents: 44568
diff changeset
   894
lemma islimpt_UNIV_iff: "x islimpt UNIV \<longleftrightarrow> \<not> open {x}"
bd91b77c4cd6 move class perfect_space into RealVector.thy;
huffman
parents: 44568
diff changeset
   895
  unfolding islimpt_def by (safe, fast, case_tac "T = {x}", fast, fast)
bd91b77c4cd6 move class perfect_space into RealVector.thy;
huffman
parents: 44568
diff changeset
   896
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   897
text {* A perfect space has no isolated points. *}
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   898
44571
bd91b77c4cd6 move class perfect_space into RealVector.thy;
huffman
parents: 44568
diff changeset
   899
lemma islimpt_UNIV [simp, intro]: "(x::'a::perfect_space) islimpt UNIV"
bd91b77c4cd6 move class perfect_space into RealVector.thy;
huffman
parents: 44568
diff changeset
   900
  unfolding islimpt_UNIV_iff by (rule not_open_singleton)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   901
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   902
lemma perfect_choose_dist:
44072
5b970711fb39 class perfect_space inherits from topological_space;
huffman
parents: 43338
diff changeset
   903
  fixes x :: "'a::{perfect_space, metric_space}"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   904
  shows "0 < r \<Longrightarrow> \<exists>a. a \<noteq> x \<and> dist a x < r"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   905
using islimpt_UNIV [of x]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   906
by (simp add: islimpt_approachable)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   907
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   908
lemma closed_limpt: "closed S \<longleftrightarrow> (\<forall>x. x islimpt S \<longrightarrow> x \<in> S)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   909
  unfolding closed_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   910
  apply (subst open_subopen)
34105
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
   911
  apply (simp add: islimpt_def subset_eq)
44170
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   912
  by (metis ComplE ComplI)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   913
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   914
lemma islimpt_EMPTY[simp]: "\<not> x islimpt {}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   915
  unfolding islimpt_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   916
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   917
lemma finite_set_avoid:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   918
  fixes a :: "'a::metric_space"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   919
  assumes fS: "finite S" shows  "\<exists>d>0. \<forall>x\<in>S. x \<noteq> a \<longrightarrow> d <= dist a x"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   920
proof(induct rule: finite_induct[OF fS])
41863
e5104b436ea1 removed dependency on Dense_Linear_Order
boehmes
parents: 41413
diff changeset
   921
  case 1 thus ?case by (auto intro: zero_less_one)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   922
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   923
  case (2 x F)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   924
  from 2 obtain d where d: "d >0" "\<forall>x\<in>F. x\<noteq>a \<longrightarrow> d \<le> dist a x" by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   925
  {assume "x = a" hence ?case using d by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   926
  moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   927
  {assume xa: "x\<noteq>a"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   928
    let ?d = "min d (dist a x)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   929
    have dp: "?d > 0" using xa d(1) using dist_nz by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   930
    from d have d': "\<forall>x\<in>F. x\<noteq>a \<longrightarrow> ?d \<le> dist a x" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   931
    with dp xa have ?case by(auto intro!: exI[where x="?d"]) }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   932
  ultimately show ?case by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   933
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   934
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   935
lemma islimpt_Un: "x islimpt (S \<union> T) \<longleftrightarrow> x islimpt S \<or> x islimpt T"
50897
078590669527 generalize lemma islimpt_finite to class t1_space
huffman
parents: 50884
diff changeset
   936
  by (simp add: islimpt_iff_eventually eventually_conj_iff)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   937
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   938
lemma discrete_imp_closed:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   939
  fixes S :: "'a::metric_space set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   940
  assumes e: "0 < e" and d: "\<forall>x \<in> S. \<forall>y \<in> S. dist y x < e \<longrightarrow> y = x"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   941
  shows "closed S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   942
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   943
  {fix x assume C: "\<forall>e>0. \<exists>x'\<in>S. x' \<noteq> x \<and> dist x' x < e"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   944
    from e have e2: "e/2 > 0" by arith
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   945
    from C[rule_format, OF e2] obtain y where y: "y \<in> S" "y\<noteq>x" "dist y x < e/2" by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   946
    let ?m = "min (e/2) (dist x y) "
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   947
    from e2 y(2) have mp: "?m > 0" by (simp add: dist_nz[THEN sym])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   948
    from C[rule_format, OF mp] obtain z where z: "z \<in> S" "z\<noteq>x" "dist z x < ?m" by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   949
    have th: "dist z y < e" using z y
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   950
      by (intro dist_triangle_lt [where z=x], simp)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   951
    from d[rule_format, OF y(1) z(1) th] y z
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   952
    have False by (auto simp add: dist_commute)}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   953
  then show ?thesis by (metis islimpt_approachable closed_limpt [where 'a='a])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   954
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   955
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   956
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   957
subsection {* Interior of a Set *}
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   958
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   959
definition "interior S = \<Union>{T. open T \<and> T \<subseteq> S}"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   960
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   961
lemma interiorI [intro?]:
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   962
  assumes "open T" and "x \<in> T" and "T \<subseteq> S"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   963
  shows "x \<in> interior S"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   964
  using assms unfolding interior_def by fast
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   965
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   966
lemma interiorE [elim?]:
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   967
  assumes "x \<in> interior S"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   968
  obtains T where "open T" and "x \<in> T" and "T \<subseteq> S"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   969
  using assms unfolding interior_def by fast
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   970
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   971
lemma open_interior [simp, intro]: "open (interior S)"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   972
  by (simp add: interior_def open_Union)
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   973
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   974
lemma interior_subset: "interior S \<subseteq> S"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   975
  by (auto simp add: interior_def)
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   976
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   977
lemma interior_maximal: "T \<subseteq> S \<Longrightarrow> open T \<Longrightarrow> T \<subseteq> interior S"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   978
  by (auto simp add: interior_def)
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   979
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   980
lemma interior_open: "open S \<Longrightarrow> interior S = S"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   981
  by (intro equalityI interior_subset interior_maximal subset_refl)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   982
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   983
lemma interior_eq: "interior S = S \<longleftrightarrow> open S"
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   984
  by (metis open_interior interior_open)
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   985
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   986
lemma open_subset_interior: "open S \<Longrightarrow> S \<subseteq> interior T \<longleftrightarrow> S \<subseteq> T"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   987
  by (metis interior_maximal interior_subset subset_trans)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   988
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   989
lemma interior_empty [simp]: "interior {} = {}"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   990
  using open_empty by (rule interior_open)
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   991
44522
2f7e9d890efe rename subset_{interior,closure} to {interior,closure}_mono;
huffman
parents: 44519
diff changeset
   992
lemma interior_UNIV [simp]: "interior UNIV = UNIV"
2f7e9d890efe rename subset_{interior,closure} to {interior,closure}_mono;
huffman
parents: 44519
diff changeset
   993
  using open_UNIV by (rule interior_open)
2f7e9d890efe rename subset_{interior,closure} to {interior,closure}_mono;
huffman
parents: 44519
diff changeset
   994
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   995
lemma interior_interior [simp]: "interior (interior S) = interior S"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   996
  using open_interior by (rule interior_open)
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   997
44522
2f7e9d890efe rename subset_{interior,closure} to {interior,closure}_mono;
huffman
parents: 44519
diff changeset
   998
lemma interior_mono: "S \<subseteq> T \<Longrightarrow> interior S \<subseteq> interior T"
2f7e9d890efe rename subset_{interior,closure} to {interior,closure}_mono;
huffman
parents: 44519
diff changeset
   999
  by (auto simp add: interior_def)
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1000
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1001
lemma interior_unique:
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1002
  assumes "T \<subseteq> S" and "open T"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1003
  assumes "\<And>T'. T' \<subseteq> S \<Longrightarrow> open T' \<Longrightarrow> T' \<subseteq> T"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1004
  shows "interior S = T"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1005
  by (intro equalityI assms interior_subset open_interior interior_maximal)
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1006
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1007
lemma interior_inter [simp]: "interior (S \<inter> T) = interior S \<inter> interior T"
44522
2f7e9d890efe rename subset_{interior,closure} to {interior,closure}_mono;
huffman
parents: 44519
diff changeset
  1008
  by (intro equalityI Int_mono Int_greatest interior_mono Int_lower1
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1009
    Int_lower2 interior_maximal interior_subset open_Int open_interior)
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1010
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1011
lemma mem_interior: "x \<in> interior S \<longleftrightarrow> (\<exists>e>0. ball x e \<subseteq> S)"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1012
  using open_contains_ball_eq [where S="interior S"]
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1013
  by (simp add: open_subset_interior)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1014
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1015
lemma interior_limit_point [intro]:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1016
  fixes x :: "'a::perfect_space"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1017
  assumes x: "x \<in> interior S" shows "x islimpt S"
44072
5b970711fb39 class perfect_space inherits from topological_space;
huffman
parents: 43338
diff changeset
  1018
  using x islimpt_UNIV [of x]
5b970711fb39 class perfect_space inherits from topological_space;
huffman
parents: 43338
diff changeset
  1019
  unfolding interior_def islimpt_def
5b970711fb39 class perfect_space inherits from topological_space;
huffman
parents: 43338
diff changeset
  1020
  apply (clarsimp, rename_tac T T')
5b970711fb39 class perfect_space inherits from topological_space;
huffman
parents: 43338
diff changeset
  1021
  apply (drule_tac x="T \<inter> T'" in spec)
5b970711fb39 class perfect_space inherits from topological_space;
huffman
parents: 43338
diff changeset
  1022
  apply (auto simp add: open_Int)
5b970711fb39 class perfect_space inherits from topological_space;
huffman
parents: 43338
diff changeset
  1023
  done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1024
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1025
lemma interior_closed_Un_empty_interior:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1026
  assumes cS: "closed S" and iT: "interior T = {}"
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1027
  shows "interior (S \<union> T) = interior S"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1028
proof
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1029
  show "interior S \<subseteq> interior (S \<union> T)"
44522
2f7e9d890efe rename subset_{interior,closure} to {interior,closure}_mono;
huffman
parents: 44519
diff changeset
  1030
    by (rule interior_mono, rule Un_upper1)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1031
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1032
  show "interior (S \<union> T) \<subseteq> interior S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1033
  proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1034
    fix x assume "x \<in> interior (S \<union> T)"
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1035
    then obtain R where "open R" "x \<in> R" "R \<subseteq> S \<union> T" ..
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1036
    show "x \<in> interior S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1037
    proof (rule ccontr)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1038
      assume "x \<notin> interior S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1039
      with `x \<in> R` `open R` obtain y where "y \<in> R - S"
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1040
        unfolding interior_def by fast
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1041
      from `open R` `closed S` have "open (R - S)" by (rule open_Diff)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1042
      from `R \<subseteq> S \<union> T` have "R - S \<subseteq> T" by fast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1043
      from `y \<in> R - S` `open (R - S)` `R - S \<subseteq> T` `interior T = {}`
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1044
      show "False" unfolding interior_def by fast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1045
    qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1046
  qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1047
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1048
44365
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1049
lemma interior_Times: "interior (A \<times> B) = interior A \<times> interior B"
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1050
proof (rule interior_unique)
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1051
  show "interior A \<times> interior B \<subseteq> A \<times> B"
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1052
    by (intro Sigma_mono interior_subset)
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1053
  show "open (interior A \<times> interior B)"
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1054
    by (intro open_Times open_interior)
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1055
  fix T assume "T \<subseteq> A \<times> B" and "open T" thus "T \<subseteq> interior A \<times> interior B"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1056
  proof (safe)
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1057
    fix x y assume "(x, y) \<in> T"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1058
    then obtain C D where "open C" "open D" "C \<times> D \<subseteq> T" "x \<in> C" "y \<in> D"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1059
      using `open T` unfolding open_prod_def by fast
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1060
    hence "open C" "open D" "C \<subseteq> A" "D \<subseteq> B" "x \<in> C" "y \<in> D"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1061
      using `T \<subseteq> A \<times> B` by auto
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1062
    thus "x \<in> interior A" and "y \<in> interior B"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1063
      by (auto intro: interiorI)
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1064
  qed
44365
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1065
qed
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1066
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1067
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1068
subsection {* Closure of a Set *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1069
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1070
definition "closure S = S \<union> {x | x. x islimpt S}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1071
44518
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
  1072
lemma interior_closure: "interior S = - (closure (- S))"
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
  1073
  unfolding interior_def closure_def islimpt_def by auto
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
  1074
34105
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
  1075
lemma closure_interior: "closure S = - interior (- S)"
44518
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
  1076
  unfolding interior_closure by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1077
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1078
lemma closed_closure[simp, intro]: "closed (closure S)"
44518
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
  1079
  unfolding closure_interior by (simp add: closed_Compl)
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
  1080
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
  1081
lemma closure_subset: "S \<subseteq> closure S"
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
  1082
  unfolding closure_def by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1083
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1084
lemma closure_hull: "closure S = closed hull S"
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1085
  unfolding hull_def closure_interior interior_def by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1086
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1087
lemma closure_eq: "closure S = S \<longleftrightarrow> closed S"
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1088
  unfolding closure_hull using closed_Inter by (rule hull_eq)
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1089
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1090
lemma closure_closed [simp]: "closed S \<Longrightarrow> closure S = S"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1091
  unfolding closure_eq .
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1092
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1093
lemma closure_closure [simp]: "closure (closure S) = closure S"
44518
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
  1094
  unfolding closure_hull by (rule hull_hull)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1095
44522
2f7e9d890efe rename subset_{interior,closure} to {interior,closure}_mono;
huffman
parents: 44519
diff changeset
  1096
lemma closure_mono: "S \<subseteq> T \<Longrightarrow> closure S \<subseteq> closure T"
44518
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
  1097
  unfolding closure_hull by (rule hull_mono)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1098
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1099
lemma closure_minimal: "S \<subseteq> T \<Longrightarrow> closed T \<Longrightarrow> closure S \<subseteq> T"
44518
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
  1100
  unfolding closure_hull by (rule hull_minimal)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1101
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1102
lemma closure_unique:
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1103
  assumes "S \<subseteq> T" and "closed T"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1104
  assumes "\<And>T'. S \<subseteq> T' \<Longrightarrow> closed T' \<Longrightarrow> T \<subseteq> T'"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1105
  shows "closure S = T"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1106
  using assms unfolding closure_hull by (rule hull_unique)
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1107
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1108
lemma closure_empty [simp]: "closure {} = {}"
44518
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
  1109
  using closed_empty by (rule closure_closed)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1110
44522
2f7e9d890efe rename subset_{interior,closure} to {interior,closure}_mono;
huffman
parents: 44519
diff changeset
  1111
lemma closure_UNIV [simp]: "closure UNIV = UNIV"
44518
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
  1112
  using closed_UNIV by (rule closure_closed)
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
  1113
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
  1114
lemma closure_union [simp]: "closure (S \<union> T) = closure S \<union> closure T"
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
  1115
  unfolding closure_interior by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1116
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1117
lemma closure_eq_empty: "closure S = {} \<longleftrightarrow> S = {}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1118
  using closure_empty closure_subset[of S]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1119
  by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1120
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1121
lemma closure_subset_eq: "closure S \<subseteq> S \<longleftrightarrow> closed S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1122
  using closure_eq[of S] closure_subset[of S]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1123
  by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1124
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1125
lemma open_inter_closure_eq_empty:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1126
  "open S \<Longrightarrow> (S \<inter> closure T) = {} \<longleftrightarrow> S \<inter> T = {}"
34105
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
  1127
  using open_subset_interior[of S "- T"]
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
  1128
  using interior_subset[of "- T"]
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1129
  unfolding closure_interior
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1130
  by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1131
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1132
lemma open_inter_closure_subset:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1133
  "open S \<Longrightarrow> (S \<inter> (closure T)) \<subseteq> closure(S \<inter> T)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1134
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1135
  fix x
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1136
  assume as: "open S" "x \<in> S \<inter> closure T"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1137
  { assume *:"x islimpt T"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1138
    have "x islimpt (S \<inter> T)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1139
    proof (rule islimptI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1140
      fix A
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1141
      assume "x \<in> A" "open A"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1142
      with as have "x \<in> A \<inter> S" "open (A \<inter> S)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1143
        by (simp_all add: open_Int)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1144
      with * obtain y where "y \<in> T" "y \<in> A \<inter> S" "y \<noteq> x"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1145
        by (rule islimptE)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1146
      hence "y \<in> S \<inter> T" "y \<in> A \<and> y \<noteq> x"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1147
        by simp_all
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1148
      thus "\<exists>y\<in>(S \<inter> T). y \<in> A \<and> y \<noteq> x" ..
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1149
    qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1150
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1151
  then show "x \<in> closure (S \<inter> T)" using as
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1152
    unfolding closure_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1153
    by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1154
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1155
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1156
lemma closure_complement: "closure (- S) = - interior S"
44518
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
  1157
  unfolding closure_interior by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1158
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1159
lemma interior_complement: "interior (- S) = - closure S"
44518
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
  1160
  unfolding closure_interior by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1161
44365
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1162
lemma closure_Times: "closure (A \<times> B) = closure A \<times> closure B"
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1163
proof (rule closure_unique)
44365
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1164
  show "A \<times> B \<subseteq> closure A \<times> closure B"
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1165
    by (intro Sigma_mono closure_subset)
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1166
  show "closed (closure A \<times> closure B)"
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1167
    by (intro closed_Times closed_closure)
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1168
  fix T assume "A \<times> B \<subseteq> T" and "closed T" thus "closure A \<times> closure B \<subseteq> T"
44365
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1169
    apply (simp add: closed_def open_prod_def, clarify)
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1170
    apply (rule ccontr)
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1171
    apply (drule_tac x="(a, b)" in bspec, simp, clarify, rename_tac C D)
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1172
    apply (simp add: closure_interior interior_def)
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1173
    apply (drule_tac x=C in spec)
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1174
    apply (drule_tac x=D in spec)
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1175
    apply auto
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1176
    done
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1177
qed
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1178
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1179
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1180
subsection {* Frontier (aka boundary) *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1181
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1182
definition "frontier S = closure S - interior S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1183
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1184
lemma frontier_closed: "closed(frontier S)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1185
  by (simp add: frontier_def closed_Diff)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1186
34105
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
  1187
lemma frontier_closures: "frontier S = (closure S) \<inter> (closure(- S))"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1188
  by (auto simp add: frontier_def interior_closure)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1189
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1190
lemma frontier_straddle:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1191
  fixes a :: "'a::metric_space"
44909
1f5d6eb73549 shorten proof of frontier_straddle
huffman
parents: 44907
diff changeset
  1192
  shows "a \<in> frontier S \<longleftrightarrow> (\<forall>e>0. (\<exists>x\<in>S. dist a x < e) \<and> (\<exists>x. x \<notin> S \<and> dist a x < e))"
1f5d6eb73549 shorten proof of frontier_straddle
huffman
parents: 44907
diff changeset
  1193
  unfolding frontier_def closure_interior
1f5d6eb73549 shorten proof of frontier_straddle
huffman
parents: 44907
diff changeset
  1194
  by (auto simp add: mem_interior subset_eq ball_def)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1195
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1196
lemma frontier_subset_closed: "closed S \<Longrightarrow> frontier S \<subseteq> S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1197
  by (metis frontier_def closure_closed Diff_subset)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1198
34964
4e8be3c04d37 Replaced vec1 and dest_vec1 by abbreviation.
hoelzl
parents: 34291
diff changeset
  1199
lemma frontier_empty[simp]: "frontier {} = {}"
36362
06475a1547cb fix lots of looping simp calls and other warnings
huffman
parents: 36360
diff changeset
  1200
  by (simp add: frontier_def)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1201
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1202
lemma frontier_subset_eq: "frontier S \<subseteq> S \<longleftrightarrow> closed S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1203
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1204
  { assume "frontier S \<subseteq> S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1205
    hence "closure S \<subseteq> S" using interior_subset unfolding frontier_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1206
    hence "closed S" using closure_subset_eq by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1207
  }
36362
06475a1547cb fix lots of looping simp calls and other warnings
huffman
parents: 36360
diff changeset
  1208
  thus ?thesis using frontier_subset_closed[of S] ..
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1209
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1210
34105
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
  1211
lemma frontier_complement: "frontier(- S) = frontier S"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1212
  by (auto simp add: frontier_def closure_complement interior_complement)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1213
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1214
lemma frontier_disjoint_eq: "frontier S \<inter> S = {} \<longleftrightarrow> open S"
34105
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
  1215
  using frontier_complement frontier_subset_eq[of "- S"]
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
  1216
  unfolding open_closed by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1217
44081
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44076
diff changeset
  1218
subsection {* Filters and the ``eventually true'' quantifier *}
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44076
diff changeset
  1219
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1220
definition
44081
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44076
diff changeset
  1221
  indirection :: "'a::real_normed_vector \<Rightarrow> 'a \<Rightarrow> 'a filter"
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44076
diff changeset
  1222
    (infixr "indirection" 70) where
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1223
  "a indirection v = (at a) within {b. \<exists>c\<ge>0. b - a = scaleR c v}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1224
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  1225
text {* Identify Trivial limits, where we can't approach arbitrarily closely. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1226
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1227
lemma trivial_limit_within:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1228
  shows "trivial_limit (at a within S) \<longleftrightarrow> \<not> a islimpt S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1229
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1230
  assume "trivial_limit (at a within S)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1231
  thus "\<not> a islimpt S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1232
    unfolding trivial_limit_def
36358
246493d61204 define nets directly as filters, instead of as filter bases
huffman
parents: 36336
diff changeset
  1233
    unfolding eventually_within eventually_at_topological
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1234
    unfolding islimpt_def
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
  1235
    apply (clarsimp simp add: set_eq_iff)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1236
    apply (rename_tac T, rule_tac x=T in exI)
36358
246493d61204 define nets directly as filters, instead of as filter bases
huffman
parents: 36336
diff changeset
  1237
    apply (clarsimp, drule_tac x=y in bspec, simp_all)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1238
    done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1239
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1240
  assume "\<not> a islimpt S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1241
  thus "trivial_limit (at a within S)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1242
    unfolding trivial_limit_def
36358
246493d61204 define nets directly as filters, instead of as filter bases
huffman
parents: 36336
diff changeset
  1243
    unfolding eventually_within eventually_at_topological
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1244
    unfolding islimpt_def
36358
246493d61204 define nets directly as filters, instead of as filter bases
huffman
parents: 36336
diff changeset
  1245
    apply clarsimp
246493d61204 define nets directly as filters, instead of as filter bases
huffman
parents: 36336
diff changeset
  1246
    apply (rule_tac x=T in exI)
246493d61204 define nets directly as filters, instead of as filter bases
huffman
parents: 36336
diff changeset
  1247
    apply auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1248
    done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1249
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1250
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1251
lemma trivial_limit_at_iff: "trivial_limit (at a) \<longleftrightarrow> \<not> a islimpt UNIV"
45031
9583f2b56f85 add lemmas within_empty and tendsto_bot;
huffman
parents: 44909
diff changeset
  1252
  using trivial_limit_within [of a UNIV] by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1253
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1254
lemma trivial_limit_at:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1255
  fixes a :: "'a::perfect_space"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1256
  shows "\<not> trivial_limit (at a)"
44571
bd91b77c4cd6 move class perfect_space into RealVector.thy;
huffman
parents: 44568
diff changeset
  1257
  by (rule at_neq_bot)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1258
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1259
lemma trivial_limit_at_infinity:
44081
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44076
diff changeset
  1260
  "\<not> trivial_limit (at_infinity :: ('a::{real_normed_vector,perfect_space}) filter)"
36358
246493d61204 define nets directly as filters, instead of as filter bases
huffman
parents: 36336
diff changeset
  1261
  unfolding trivial_limit_def eventually_at_infinity
246493d61204 define nets directly as filters, instead of as filter bases
huffman
parents: 36336
diff changeset
  1262
  apply clarsimp
44072
5b970711fb39 class perfect_space inherits from topological_space;
huffman
parents: 43338
diff changeset
  1263
  apply (subgoal_tac "\<exists>x::'a. x \<noteq> 0", clarify)
5b970711fb39 class perfect_space inherits from topological_space;
huffman
parents: 43338
diff changeset
  1264
   apply (rule_tac x="scaleR (b / norm x) x" in exI, simp)
5b970711fb39 class perfect_space inherits from topological_space;
huffman
parents: 43338
diff changeset
  1265
  apply (cut_tac islimpt_UNIV [of "0::'a", unfolded islimpt_def])
5b970711fb39 class perfect_space inherits from topological_space;
huffman
parents: 43338
diff changeset
  1266
  apply (drule_tac x=UNIV in spec, simp)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1267
  done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1268
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  1269
text {* Some property holds "sufficiently close" to the limit point. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1270
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1271
lemma eventually_at: (* FIXME: this replaces Limits.eventually_at *)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1272
  "eventually P (at a) \<longleftrightarrow> (\<exists>d>0. \<forall>x. 0 < dist x a \<and> dist x a < d \<longrightarrow> P x)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1273
unfolding eventually_at dist_nz by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1274
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  1275
lemma eventually_within: (* FIXME: this replaces Limits.eventually_within *)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  1276
  "eventually P (at a within S) \<longleftrightarrow>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1277
        (\<exists>d>0. \<forall>x\<in>S. 0 < dist x a \<and> dist x a < d \<longrightarrow> P x)"
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  1278
  by (rule eventually_within_less)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1279
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1280
lemma eventually_happens: "eventually P net ==> trivial_limit net \<or> (\<exists>x. P x)"
36358
246493d61204 define nets directly as filters, instead of as filter bases
huffman
parents: 36336
diff changeset
  1281
  unfolding trivial_limit_def
246493d61204 define nets directly as filters, instead of as filter bases
huffman
parents: 36336
diff changeset
  1282
  by (auto elim: eventually_rev_mp)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1283
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1284
lemma trivial_limit_eventually: "trivial_limit net \<Longrightarrow> eventually P net"
45031
9583f2b56f85 add lemmas within_empty and tendsto_bot;
huffman
parents: 44909
diff changeset
  1285
  by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1286
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1287
lemma trivial_limit_eq: "trivial_limit net \<longleftrightarrow> (\<forall>P. eventually P net)"
44342
8321948340ea redefine constant 'trivial_limit' as an abbreviation
huffman
parents: 44286
diff changeset
  1288
  by (simp add: filter_eq_iff)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1289
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1290
text{* Combining theorems for "eventually" *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1291
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1292
lemma eventually_rev_mono:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1293
  "eventually P net \<Longrightarrow> (\<forall>x. P x \<longrightarrow> Q x) \<Longrightarrow> eventually Q net"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1294
using eventually_mono [of P Q] by fast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1295
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1296
lemma not_eventually: "(\<forall>x. \<not> P x ) \<Longrightarrow> ~(trivial_limit net) ==> ~(eventually (\<lambda>x. P x) net)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1297
  by (simp add: eventually_False)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1298
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1299
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  1300
subsection {* Limits *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1301
44081
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44076
diff changeset
  1302
text{* Notation Lim to avoid collition with lim defined in analysis *}
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44076
diff changeset
  1303
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44076
diff changeset
  1304
definition Lim :: "'a filter \<Rightarrow> ('a \<Rightarrow> 'b::t2_space) \<Rightarrow> 'b"
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44076
diff changeset
  1305
  where "Lim A f = (THE l. (f ---> l) A)"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1306
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1307
lemma Lim:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1308
 "(f ---> l) net \<longleftrightarrow>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1309
        trivial_limit net \<or>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1310
        (\<forall>e>0. eventually (\<lambda>x. dist (f x) l < e) net)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1311
  unfolding tendsto_iff trivial_limit_eq by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1312
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1313
text{* Show that they yield usual definitions in the various cases. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1314
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1315
lemma Lim_within_le: "(f ---> l)(at a within S) \<longleftrightarrow>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1316
           (\<forall>e>0. \<exists>d>0. \<forall>x\<in>S. 0 < dist x a  \<and> dist x a  <= d \<longrightarrow> dist (f x) l < e)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1317
  by (auto simp add: tendsto_iff eventually_within_le)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1318
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1319
lemma Lim_within: "(f ---> l) (at a within S) \<longleftrightarrow>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1320
        (\<forall>e >0. \<exists>d>0. \<forall>x \<in> S. 0 < dist x a  \<and> dist x a  < d  \<longrightarrow> dist (f x) l < e)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1321
  by (auto simp add: tendsto_iff eventually_within)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1322
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1323
lemma Lim_at: "(f ---> l) (at a) \<longleftrightarrow>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1324
        (\<forall>e >0. \<exists>d>0. \<forall>x. 0 < dist x a  \<and> dist x a  < d  \<longrightarrow> dist (f x) l < e)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1325
  by (auto simp add: tendsto_iff eventually_at)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1326
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1327
lemma Lim_at_infinity:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1328
  "(f ---> l) at_infinity \<longleftrightarrow> (\<forall>e>0. \<exists>b. \<forall>x. norm x >= b \<longrightarrow> dist (f x) l < e)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1329
  by (auto simp add: tendsto_iff eventually_at_infinity)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1330
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1331
lemma Lim_eventually: "eventually (\<lambda>x. f x = l) net \<Longrightarrow> (f ---> l) net"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1332
  by (rule topological_tendstoI, auto elim: eventually_rev_mono)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1333
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1334
text{* The expected monotonicity property. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1335
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1336
lemma Lim_within_empty: "(f ---> l) (net within {})"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1337
  unfolding tendsto_def Limits.eventually_within by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1338
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1339
lemma Lim_within_subset: "(f ---> l) (net within S) \<Longrightarrow> T \<subseteq> S \<Longrightarrow> (f ---> l) (net within T)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1340
  unfolding tendsto_def Limits.eventually_within
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1341
  by (auto elim!: eventually_elim1)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1342
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1343
lemma Lim_Un: assumes "(f ---> l) (net within S)" "(f ---> l) (net within T)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1344
  shows "(f ---> l) (net within (S \<union> T))"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1345
  using assms unfolding tendsto_def Limits.eventually_within
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1346
  apply clarify
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1347
  apply (drule spec, drule (1) mp, drule (1) mp)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1348
  apply (drule spec, drule (1) mp, drule (1) mp)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1349
  apply (auto elim: eventually_elim2)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1350
  done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1351
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1352
lemma Lim_Un_univ:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1353
 "(f ---> l) (net within S) \<Longrightarrow> (f ---> l) (net within T) \<Longrightarrow>  S \<union> T = UNIV
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1354
        ==> (f ---> l) net"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1355
  by (metis Lim_Un within_UNIV)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1356
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1357
text{* Interrelations between restricted and unrestricted limits. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1358
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1359
lemma Lim_at_within: "(f ---> l) net ==> (f ---> l)(net within S)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1360
  (* FIXME: rename *)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1361
  unfolding tendsto_def Limits.eventually_within
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1362
  apply (clarify, drule spec, drule (1) mp, drule (1) mp)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1363
  by (auto elim!: eventually_elim1)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1364
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1365
lemma eventually_within_interior:
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1366
  assumes "x \<in> interior S"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1367
  shows "eventually P (at x within S) \<longleftrightarrow> eventually P (at x)" (is "?lhs = ?rhs")
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1368
proof-
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1369
  from assms obtain T where T: "open T" "x \<in> T" "T \<subseteq> S" ..
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1370
  { assume "?lhs"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1371
    then obtain A where "open A" "x \<in> A" "\<forall>y\<in>A. y \<noteq> x \<longrightarrow> y \<in> S \<longrightarrow> P y"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1372
      unfolding Limits.eventually_within Limits.eventually_at_topological
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1373
      by auto
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1374
    with T have "open (A \<inter> T)" "x \<in> A \<inter> T" "\<forall>y\<in>(A \<inter> T). y \<noteq> x \<longrightarrow> P y"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1375
      by auto
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1376
    then have "?rhs"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1377
      unfolding Limits.eventually_at_topological by auto
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1378
  } moreover
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1379
  { assume "?rhs" hence "?lhs"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1380
      unfolding Limits.eventually_within
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1381
      by (auto elim: eventually_elim1)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1382
  } ultimately
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1383
  show "?thesis" ..
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1384
qed
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1385
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1386
lemma at_within_interior:
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1387
  "x \<in> interior S \<Longrightarrow> at x within S = at x"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1388
  by (simp add: filter_eq_iff eventually_within_interior)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1389
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1390
lemma at_within_open:
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1391
  "\<lbrakk>x \<in> S; open S\<rbrakk> \<Longrightarrow> at x within S = at x"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1392
  by (simp only: at_within_interior interior_open)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1393
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1394
lemma Lim_within_open:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1395
  fixes f :: "'a::topological_space \<Rightarrow> 'b::topological_space"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1396
  assumes"a \<in> S" "open S"
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1397
  shows "(f ---> l)(at a within S) \<longleftrightarrow> (f ---> l)(at a)"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1398
  using assms by (simp only: at_within_open)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1399
43338
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1400
lemma Lim_within_LIMSEQ:
44584
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  1401
  fixes a :: "'a::metric_space"
43338
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1402
  assumes "\<forall>S. (\<forall>n. S n \<noteq> a \<and> S n \<in> T) \<and> S ----> a \<longrightarrow> (\<lambda>n. X (S n)) ----> L"
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1403
  shows "(X ---> L) (at a within T)"
44584
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  1404
  using assms unfolding tendsto_def [where l=L]
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  1405
  by (simp add: sequentially_imp_eventually_within)
43338
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1406
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1407
lemma Lim_right_bound:
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1408
  fixes f :: "real \<Rightarrow> real"
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1409
  assumes mono: "\<And>a b. a \<in> I \<Longrightarrow> b \<in> I \<Longrightarrow> x < a \<Longrightarrow> a \<le> b \<Longrightarrow> f a \<le> f b"
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1410
  assumes bnd: "\<And>a. a \<in> I \<Longrightarrow> x < a \<Longrightarrow> K \<le> f a"
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1411
  shows "(f ---> Inf (f ` ({x<..} \<inter> I))) (at x within ({x<..} \<inter> I))"
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1412
proof cases
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1413
  assume "{x<..} \<inter> I = {}" then show ?thesis by (simp add: Lim_within_empty)
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1414
next
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1415
  assume [simp]: "{x<..} \<inter> I \<noteq> {}"
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1416
  show ?thesis
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1417
  proof (rule Lim_within_LIMSEQ, safe)
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1418
    fix S assume S: "\<forall>n. S n \<noteq> x \<and> S n \<in> {x <..} \<inter> I" "S ----> x"
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1419
    
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1420
    show "(\<lambda>n. f (S n)) ----> Inf (f ` ({x<..} \<inter> I))"
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1421
    proof (rule LIMSEQ_I, rule ccontr)
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1422
      fix r :: real assume "0 < r"
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1423
      with Inf_close[of "f ` ({x<..} \<inter> I)" r]
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1424
      obtain y where y: "x < y" "y \<in> I" "f y < Inf (f ` ({x <..} \<inter> I)) + r" by auto
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1425
      from `x < y` have "0 < y - x" by auto
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1426
      from S(2)[THEN LIMSEQ_D, OF this]
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1427
      obtain N where N: "\<And>n. N \<le> n \<Longrightarrow> \<bar>S n - x\<bar> < y - x" by auto
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1428
      
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1429
      assume "\<not> (\<exists>N. \<forall>n\<ge>N. norm (f (S n) - Inf (f ` ({x<..} \<inter> I))) < r)"
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1430
      moreover have "\<And>n. Inf (f ` ({x<..} \<inter> I)) \<le> f (S n)"
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1431
        using S bnd by (intro Inf_lower[where z=K]) auto
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1432
      ultimately obtain n where n: "N \<le> n" "r + Inf (f ` ({x<..} \<inter> I)) \<le> f (S n)"
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1433
        by (auto simp: not_less field_simps)
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1434
      with N[OF n(1)] mono[OF _ `y \<in> I`, of "S n"] S(1)[THEN spec, of n] y
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1435
      show False by auto
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1436
    qed
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1437
  qed
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1438
qed
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1439
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1440
text{* Another limit point characterization. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1441
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1442
lemma islimpt_sequential:
50883
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1443
  fixes x :: "'a::first_countable_topology"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1444
  shows "x islimpt S \<longleftrightarrow> (\<exists>f. (\<forall>n::nat. f n \<in> S - {x}) \<and> (f ---> x) sequentially)"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1445
    (is "?lhs = ?rhs")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1446
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1447
  assume ?lhs
50883
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1448
  from countable_basis_at_decseq[of x] guess A . note A = this
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1449
  def f \<equiv> "\<lambda>n. SOME y. y \<in> S \<and> y \<in> A n \<and> x \<noteq> y"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1450
  { fix n
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1451
    from `?lhs` have "\<exists>y. y \<in> S \<and> y \<in> A n \<and> x \<noteq> y"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1452
      unfolding islimpt_def using A(1,2)[of n] by auto
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1453
    then have "f n \<in> S \<and> f n \<in> A n \<and> x \<noteq> f n"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1454
      unfolding f_def by (rule someI_ex)
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1455
    then have "f n \<in> S" "f n \<in> A n" "x \<noteq> f n" by auto }
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1456
  then have "\<forall>n. f n \<in> S - {x}" by auto
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1457
  moreover have "(\<lambda>n. f n) ----> x"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1458
  proof (rule topological_tendstoI)
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1459
    fix S assume "open S" "x \<in> S"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1460
    from A(3)[OF this] `\<And>n. f n \<in> A n`
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1461
    show "eventually (\<lambda>x. f x \<in> S) sequentially" by (auto elim!: eventually_elim1)
44584
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  1462
  qed
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  1463
  ultimately show ?rhs by fast
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1464
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1465
  assume ?rhs
50883
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1466
  then obtain f :: "nat \<Rightarrow> 'a" where f: "\<And>n. f n \<in> S - {x}" and lim: "f ----> x" by auto
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1467
  show ?lhs
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1468
    unfolding islimpt_def
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1469
  proof safe
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1470
    fix T assume "open T" "x \<in> T"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1471
    from lim[THEN topological_tendstoD, OF this] f
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1472
    show "\<exists>y\<in>S. y \<in> T \<and> y \<noteq> x"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1473
      unfolding eventually_sequentially by auto
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1474
  qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1475
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1476
44125
230a8665c919 mark some redundant theorems as legacy
huffman
parents: 44122
diff changeset
  1477
lemma Lim_inv: (* TODO: delete *)
44081
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44076
diff changeset
  1478
  fixes f :: "'a \<Rightarrow> real" and A :: "'a filter"
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44076
diff changeset
  1479
  assumes "(f ---> l) A" and "l \<noteq> 0"
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44076
diff changeset
  1480
  shows "((inverse o f) ---> inverse l) A"
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  1481
  unfolding o_def using assms by (rule tendsto_inverse)
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  1482
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1483
lemma Lim_null:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1484
  fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
44125
230a8665c919 mark some redundant theorems as legacy
huffman
parents: 44122
diff changeset
  1485
  shows "(f ---> l) net \<longleftrightarrow> ((\<lambda>x. f(x) - l) ---> 0) net"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1486
  by (simp add: Lim dist_norm)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1487
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1488
lemma Lim_null_comparison:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1489
  fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1490
  assumes "eventually (\<lambda>x. norm (f x) \<le> g x) net" "(g ---> 0) net"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1491
  shows "(f ---> 0) net"
44252
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1492
proof (rule metric_tendsto_imp_tendsto)
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1493
  show "(g ---> 0) net" by fact
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1494
  show "eventually (\<lambda>x. dist (f x) 0 \<le> dist (g x) 0) net"
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1495
    using assms(1) by (rule eventually_elim1, simp add: dist_norm)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1496
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1497
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1498
lemma Lim_transform_bound:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1499
  fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1500
  fixes g :: "'a \<Rightarrow> 'c::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1501
  assumes "eventually (\<lambda>n. norm(f n) <= norm(g n)) net"  "(g ---> 0) net"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1502
  shows "(f ---> 0) net"
44252
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1503
  using assms(1) tendsto_norm_zero [OF assms(2)]
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1504
  by (rule Lim_null_comparison)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1505
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1506
text{* Deducing things about the limit from the elements. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1507
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1508
lemma Lim_in_closed_set:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1509
  assumes "closed S" "eventually (\<lambda>x. f(x) \<in> S) net" "\<not>(trivial_limit net)" "(f ---> l) net"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1510
  shows "l \<in> S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1511
proof (rule ccontr)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1512
  assume "l \<notin> S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1513
  with `closed S` have "open (- S)" "l \<in> - S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1514
    by (simp_all add: open_Compl)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1515
  with assms(4) have "eventually (\<lambda>x. f x \<in> - S) net"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1516
    by (rule topological_tendstoD)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1517
  with assms(2) have "eventually (\<lambda>x. False) net"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1518
    by (rule eventually_elim2) simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1519
  with assms(3) show "False"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1520
    by (simp add: eventually_False)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1521
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1522
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1523
text{* Need to prove closed(cball(x,e)) before deducing this as a corollary. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1524
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1525
lemma Lim_dist_ubound:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1526
  assumes "\<not>(trivial_limit net)" "(f ---> l) net" "eventually (\<lambda>x. dist a (f x) <= e) net"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1527
  shows "dist a l <= e"
44252
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1528
proof-
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1529
  have "dist a l \<in> {..e}"
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1530
  proof (rule Lim_in_closed_set)
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1531
    show "closed {..e}" by simp
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1532
    show "eventually (\<lambda>x. dist a (f x) \<in> {..e}) net" by (simp add: assms)
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1533
    show "\<not> trivial_limit net" by fact
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1534
    show "((\<lambda>x. dist a (f x)) ---> dist a l) net" by (intro tendsto_intros assms)
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1535
  qed
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1536
  thus ?thesis by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1537
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1538
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1539
lemma Lim_norm_ubound:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1540
  fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1541
  assumes "\<not>(trivial_limit net)" "(f ---> l) net" "eventually (\<lambda>x. norm(f x) <= e) net"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1542
  shows "norm(l) <= e"
44252
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1543
proof-
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1544
  have "norm l \<in> {..e}"
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1545
  proof (rule Lim_in_closed_set)
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1546
    show "closed {..e}" by simp
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1547
    show "eventually (\<lambda>x. norm (f x) \<in> {..e}) net" by (simp add: assms)
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1548
    show "\<not> trivial_limit net" by fact
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1549
    show "((\<lambda>x. norm (f x)) ---> norm l) net" by (intro tendsto_intros assms)
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1550
  qed
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1551
  thus ?thesis by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1552
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1553
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1554
lemma Lim_norm_lbound:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1555
  fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1556
  assumes "\<not> (trivial_limit net)"  "(f ---> l) net"  "eventually (\<lambda>x. e <= norm(f x)) net"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1557
  shows "e \<le> norm l"
44252
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1558
proof-
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1559
  have "norm l \<in> {e..}"
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1560
  proof (rule Lim_in_closed_set)
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1561
    show "closed {e..}" by simp
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1562
    show "eventually (\<lambda>x. norm (f x) \<in> {e..}) net" by (simp add: assms)
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1563
    show "\<not> trivial_limit net" by fact
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1564
    show "((\<lambda>x. norm (f x)) ---> norm l) net" by (intro tendsto_intros assms)
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1565
  qed
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1566
  thus ?thesis by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1567
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1568
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1569
text{* Uniqueness of the limit, when nontrivial. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1570
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1571
lemma tendsto_Lim:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1572
  fixes f :: "'a \<Rightarrow> 'b::t2_space"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1573
  shows "~(trivial_limit net) \<Longrightarrow> (f ---> l) net ==> Lim net f = l"
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41969
diff changeset
  1574
  unfolding Lim_def using tendsto_unique[of net f] by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1575
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1576
text{* Limit under bilinear function *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1577
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1578
lemma Lim_bilinear:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1579
  assumes "(f ---> l) net" and "(g ---> m) net" and "bounded_bilinear h"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1580
  shows "((\<lambda>x. h (f x) (g x)) ---> (h l m)) net"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1581
using `bounded_bilinear h` `(f ---> l) net` `(g ---> m) net`
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1582
by (rule bounded_bilinear.tendsto)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1583
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1584
text{* These are special for limits out of the same vector space. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1585
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1586
lemma Lim_within_id: "(id ---> a) (at a within s)"
45031
9583f2b56f85 add lemmas within_empty and tendsto_bot;
huffman
parents: 44909
diff changeset
  1587
  unfolding id_def by (rule tendsto_ident_at_within)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1588
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1589
lemma Lim_at_id: "(id ---> a) (at a)"
45031
9583f2b56f85 add lemmas within_empty and tendsto_bot;
huffman
parents: 44909
diff changeset
  1590
  unfolding id_def by (rule tendsto_ident_at)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1591
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1592
lemma Lim_at_zero:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1593
  fixes a :: "'a::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1594
  fixes l :: "'b::topological_space"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1595
  shows "(f ---> l) (at a) \<longleftrightarrow> ((\<lambda>x. f(a + x)) ---> l) (at 0)" (is "?lhs = ?rhs")
44252
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1596
  using LIM_offset_zero LIM_offset_zero_cancel ..
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1597
44081
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44076
diff changeset
  1598
text{* It's also sometimes useful to extract the limit point from the filter. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1599
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1600
definition
44081
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44076
diff changeset
  1601
  netlimit :: "'a::t2_space filter \<Rightarrow> 'a" where
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1602
  "netlimit net = (SOME a. ((\<lambda>x. x) ---> a) net)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1603
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1604
lemma netlimit_within:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1605
  assumes "\<not> trivial_limit (at a within S)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1606
  shows "netlimit (at a within S) = a"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1607
unfolding netlimit_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1608
apply (rule some_equality)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1609
apply (rule Lim_at_within)
44568
e6f291cb5810 discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
huffman
parents: 44533
diff changeset
  1610
apply (rule tendsto_ident_at)
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41969
diff changeset
  1611
apply (erule tendsto_unique [OF assms])
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1612
apply (rule Lim_at_within)
44568
e6f291cb5810 discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
huffman
parents: 44533
diff changeset
  1613
apply (rule tendsto_ident_at)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1614
done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1615
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1616
lemma netlimit_at:
44072
5b970711fb39 class perfect_space inherits from topological_space;
huffman
parents: 43338
diff changeset
  1617
  fixes a :: "'a::{perfect_space,t2_space}"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1618
  shows "netlimit (at a) = a"
45031
9583f2b56f85 add lemmas within_empty and tendsto_bot;
huffman
parents: 44909
diff changeset
  1619
  using netlimit_within [of a UNIV] by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1620
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1621
lemma lim_within_interior:
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1622
  "x \<in> interior S \<Longrightarrow> (f ---> l) (at x within S) \<longleftrightarrow> (f ---> l) (at x)"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1623
  by (simp add: at_within_interior)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1624
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1625
lemma netlimit_within_interior:
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1626
  fixes x :: "'a::{t2_space,perfect_space}"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1627
  assumes "x \<in> interior S"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1628
  shows "netlimit (at x within S) = x"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1629
using assms by (simp add: at_within_interior netlimit_at)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1630
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1631
text{* Transformation of limit. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1632
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1633
lemma Lim_transform:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1634
  fixes f g :: "'a::type \<Rightarrow> 'b::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1635
  assumes "((\<lambda>x. f x - g x) ---> 0) net" "(f ---> l) net"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1636
  shows "(g ---> l) net"
44252
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1637
  using tendsto_diff [OF assms(2) assms(1)] by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1638
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1639
lemma Lim_transform_eventually:
36667
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1640
  "eventually (\<lambda>x. f x = g x) net \<Longrightarrow> (f ---> l) net \<Longrightarrow> (g ---> l) net"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1641
  apply (rule topological_tendstoI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1642
  apply (drule (2) topological_tendstoD)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1643
  apply (erule (1) eventually_elim2, simp)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1644
  done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1645
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1646
lemma Lim_transform_within:
36667
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1647
  assumes "0 < d" and "\<forall>x'\<in>S. 0 < dist x' x \<and> dist x' x < d \<longrightarrow> f x' = g x'"
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1648
  and "(f ---> l) (at x within S)"
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1649
  shows "(g ---> l) (at x within S)"
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1650
proof (rule Lim_transform_eventually)
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1651
  show "eventually (\<lambda>x. f x = g x) (at x within S)"
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1652
    unfolding eventually_within
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1653
    using assms(1,2) by auto
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1654
  show "(f ---> l) (at x within S)" by fact
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1655
qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1656
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1657
lemma Lim_transform_at:
36667
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1658
  assumes "0 < d" and "\<forall>x'. 0 < dist x' x \<and> dist x' x < d \<longrightarrow> f x' = g x'"
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1659
  and "(f ---> l) (at x)"
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1660
  shows "(g ---> l) (at x)"
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1661
proof (rule Lim_transform_eventually)
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1662
  show "eventually (\<lambda>x. f x = g x) (at x)"
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1663
    unfolding eventually_at
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1664
    using assms(1,2) by auto
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1665
  show "(f ---> l) (at x)" by fact
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1666
qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1667
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1668
text{* Common case assuming being away from some crucial point like 0. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1669
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1670
lemma Lim_transform_away_within:
36669
c90c8a3ae1f7 generalize some lemmas to class t1_space
huffman
parents: 36668
diff changeset
  1671
  fixes a b :: "'a::t1_space"
36667
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1672
  assumes "a \<noteq> b" and "\<forall>x\<in>S. x \<noteq> a \<and> x \<noteq> b \<longrightarrow> f x = g x"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1673
  and "(f ---> l) (at a within S)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1674
  shows "(g ---> l) (at a within S)"
36669
c90c8a3ae1f7 generalize some lemmas to class t1_space
huffman
parents: 36668
diff changeset
  1675
proof (rule Lim_transform_eventually)
c90c8a3ae1f7 generalize some lemmas to class t1_space
huffman
parents: 36668
diff changeset
  1676
  show "(f ---> l) (at a within S)" by fact
c90c8a3ae1f7 generalize some lemmas to class t1_space
huffman
parents: 36668
diff changeset
  1677
  show "eventually (\<lambda>x. f x = g x) (at a within S)"
c90c8a3ae1f7 generalize some lemmas to class t1_space
huffman
parents: 36668
diff changeset
  1678
    unfolding Limits.eventually_within eventually_at_topological
c90c8a3ae1f7 generalize some lemmas to class t1_space
huffman
parents: 36668
diff changeset
  1679
    by (rule exI [where x="- {b}"], simp add: open_Compl assms)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1680
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1681
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1682
lemma Lim_transform_away_at:
36669
c90c8a3ae1f7 generalize some lemmas to class t1_space
huffman
parents: 36668
diff changeset
  1683
  fixes a b :: "'a::t1_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1684
  assumes ab: "a\<noteq>b" and fg: "\<forall>x. x \<noteq> a \<and> x \<noteq> b \<longrightarrow> f x = g x"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1685
  and fl: "(f ---> l) (at a)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1686
  shows "(g ---> l) (at a)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1687
  using Lim_transform_away_within[OF ab, of UNIV f g l] fg fl
45031
9583f2b56f85 add lemmas within_empty and tendsto_bot;
huffman
parents: 44909
diff changeset
  1688
  by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1689
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1690
text{* Alternatively, within an open set. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1691
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1692
lemma Lim_transform_within_open:
36667
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1693
  assumes "open S" and "a \<in> S" and "\<forall>x\<in>S. x \<noteq> a \<longrightarrow> f x = g x"
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1694
  and "(f ---> l) (at a)"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1695
  shows "(g ---> l) (at a)"
36667
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1696
proof (rule Lim_transform_eventually)
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1697
  show "eventually (\<lambda>x. f x = g x) (at a)"
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1698
    unfolding eventually_at_topological
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1699
    using assms(1,2,3) by auto
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1700
  show "(f ---> l) (at a)" by fact
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1701
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1702
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1703
text{* A congruence rule allowing us to transform limits assuming not at point. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1704
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1705
(* FIXME: Only one congruence rule for tendsto can be used at a time! *)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1706
36362
06475a1547cb fix lots of looping simp calls and other warnings
huffman
parents: 36360
diff changeset
  1707
lemma Lim_cong_within(*[cong add]*):
43338
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1708
  assumes "a = b" "x = y" "S = T"
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1709
  assumes "\<And>x. x \<noteq> b \<Longrightarrow> x \<in> T \<Longrightarrow> f x = g x"
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1710
  shows "(f ---> x) (at a within S) \<longleftrightarrow> (g ---> y) (at b within T)"
36667
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1711
  unfolding tendsto_def Limits.eventually_within eventually_at_topological
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1712
  using assms by simp
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1713
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1714
lemma Lim_cong_at(*[cong add]*):
43338
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1715
  assumes "a = b" "x = y"
36667
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1716
  assumes "\<And>x. x \<noteq> a \<Longrightarrow> f x = g x"
43338
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1717
  shows "((\<lambda>x. f x) ---> x) (at a) \<longleftrightarrow> ((g ---> y) (at a))"
36667
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1718
  unfolding tendsto_def eventually_at_topological
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1719
  using assms by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1720
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1721
text{* Useful lemmas on closure and set of possible sequential limits.*}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1722
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1723
lemma closure_sequential:
50883
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1724
  fixes l :: "'a::first_countable_topology"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1725
  shows "l \<in> closure S \<longleftrightarrow> (\<exists>x. (\<forall>n. x n \<in> S) \<and> (x ---> l) sequentially)" (is "?lhs = ?rhs")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1726
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1727
  assume "?lhs" moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1728
  { assume "l \<in> S"
44125
230a8665c919 mark some redundant theorems as legacy
huffman
parents: 44122
diff changeset
  1729
    hence "?rhs" using tendsto_const[of l sequentially] by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1730
  } moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1731
  { assume "l islimpt S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1732
    hence "?rhs" unfolding islimpt_sequential by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1733
  } ultimately
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1734
  show "?rhs" unfolding closure_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1735
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1736
  assume "?rhs"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1737
  thus "?lhs" unfolding closure_def unfolding islimpt_sequential by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1738
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1739
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1740
lemma closed_sequential_limits:
50883
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1741
  fixes S :: "'a::first_countable_topology set"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1742
  shows "closed S \<longleftrightarrow> (\<forall>x l. (\<forall>n. x n \<in> S) \<and> (x ---> l) sequentially \<longrightarrow> l \<in> S)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1743
  unfolding closed_limpt
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1744
  using closure_sequential [where 'a='a] closure_closed [where 'a='a] closed_limpt [where 'a='a] islimpt_sequential [where 'a='a] mem_delete [where 'a='a]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1745
  by metis
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1746
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1747
lemma closure_approachable:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1748
  fixes S :: "'a::metric_space set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1749
  shows "x \<in> closure S \<longleftrightarrow> (\<forall>e>0. \<exists>y\<in>S. dist y x < e)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1750
  apply (auto simp add: closure_def islimpt_approachable)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1751
  by (metis dist_self)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1752
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1753
lemma closed_approachable:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1754
  fixes S :: "'a::metric_space set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1755
  shows "closed S ==> (\<forall>e>0. \<exists>y\<in>S. dist y x < e) \<longleftrightarrow> x \<in> S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1756
  by (metis closure_closed closure_approachable)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1757
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1758
subsection {* Infimum Distance *}
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1759
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1760
definition "infdist x A = (if A = {} then 0 else Inf {dist x a|a. a \<in> A})"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1761
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1762
lemma infdist_notempty: "A \<noteq> {} \<Longrightarrow> infdist x A = Inf {dist x a|a. a \<in> A}"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1763
  by (simp add: infdist_def)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1764
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1765
lemma infdist_nonneg:
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1766
  shows "0 \<le> infdist x A"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1767
  using assms by (auto simp add: infdist_def)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1768
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1769
lemma infdist_le:
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1770
  assumes "a \<in> A"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1771
  assumes "d = dist x a"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1772
  shows "infdist x A \<le> d"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1773
  using assms by (auto intro!: SupInf.Inf_lower[where z=0] simp add: infdist_def)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1774
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1775
lemma infdist_zero[simp]:
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1776
  assumes "a \<in> A" shows "infdist a A = 0"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1777
proof -
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1778
  from infdist_le[OF assms, of "dist a a"] have "infdist a A \<le> 0" by auto
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1779
  with infdist_nonneg[of a A] assms show "infdist a A = 0" by auto
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1780
qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1781
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1782
lemma infdist_triangle:
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1783
  shows "infdist x A \<le> infdist y A + dist x y"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1784
proof cases
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1785
  assume "A = {}" thus ?thesis by (simp add: infdist_def)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1786
next
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1787
  assume "A \<noteq> {}" then obtain a where "a \<in> A" by auto
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1788
  have "infdist x A \<le> Inf {dist x y + dist y a |a. a \<in> A}"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1789
  proof
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1790
    from `A \<noteq> {}` show "{dist x y + dist y a |a. a \<in> A} \<noteq> {}" by simp
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1791
    fix d assume "d \<in> {dist x y + dist y a |a. a \<in> A}"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1792
    then obtain a where d: "d = dist x y + dist y a" "a \<in> A" by auto
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1793
    show "infdist x A \<le> d"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1794
      unfolding infdist_notempty[OF `A \<noteq> {}`]
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1795
    proof (rule Inf_lower2)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1796
      show "dist x a \<in> {dist x a |a. a \<in> A}" using `a \<in> A` by auto
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1797
      show "dist x a \<le> d" unfolding d by (rule dist_triangle)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1798
      fix d assume "d \<in> {dist x a |a. a \<in> A}"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1799
      then obtain a where "a \<in> A" "d = dist x a" by auto
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1800
      thus "infdist x A \<le> d" by (rule infdist_le)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1801
    qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1802
  qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1803
  also have "\<dots> = dist x y + infdist y A"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1804
  proof (rule Inf_eq, safe)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1805
    fix a assume "a \<in> A"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1806
    thus "dist x y + infdist y A \<le> dist x y + dist y a" by (auto intro: infdist_le)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1807
  next
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1808
    fix i assume inf: "\<And>d. d \<in> {dist x y + dist y a |a. a \<in> A} \<Longrightarrow> i \<le> d"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1809
    hence "i - dist x y \<le> infdist y A" unfolding infdist_notempty[OF `A \<noteq> {}`] using `a \<in> A`
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1810
      by (intro Inf_greatest) (auto simp: field_simps)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1811
    thus "i \<le> dist x y + infdist y A" by simp
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1812
  qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1813
  finally show ?thesis by simp
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1814
qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1815
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1816
lemma
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1817
  in_closure_iff_infdist_zero:
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1818
  assumes "A \<noteq> {}"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1819
  shows "x \<in> closure A \<longleftrightarrow> infdist x A = 0"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1820
proof
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1821
  assume "x \<in> closure A"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1822
  show "infdist x A = 0"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1823
  proof (rule ccontr)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1824
    assume "infdist x A \<noteq> 0"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1825
    with infdist_nonneg[of x A] have "infdist x A > 0" by auto
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1826
    hence "ball x (infdist x A) \<inter> closure A = {}" apply auto
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1827
      by (metis `0 < infdist x A` `x \<in> closure A` closure_approachable dist_commute
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1828
        eucl_less_not_refl euclidean_trans(2) infdist_le)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1829
    hence "x \<notin> closure A" by (metis `0 < infdist x A` centre_in_ball disjoint_iff_not_equal)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1830
    thus False using `x \<in> closure A` by simp
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1831
  qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1832
next
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1833
  assume x: "infdist x A = 0"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1834
  then obtain a where "a \<in> A" by atomize_elim (metis all_not_in_conv assms)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1835
  show "x \<in> closure A" unfolding closure_approachable
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1836
  proof (safe, rule ccontr)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1837
    fix e::real assume "0 < e"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1838
    assume "\<not> (\<exists>y\<in>A. dist y x < e)"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1839
    hence "infdist x A \<ge> e" using `a \<in> A`
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1840
      unfolding infdist_def
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  1841
      by (force simp: dist_commute)
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1842
    with x `0 < e` show False by auto
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1843
  qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1844
qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1845
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1846
lemma
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1847
  in_closed_iff_infdist_zero:
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1848
  assumes "closed A" "A \<noteq> {}"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1849
  shows "x \<in> A \<longleftrightarrow> infdist x A = 0"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1850
proof -
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1851
  have "x \<in> closure A \<longleftrightarrow> infdist x A = 0"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1852
    by (rule in_closure_iff_infdist_zero) fact
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1853
  with assms show ?thesis by simp
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1854
qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1855
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1856
lemma tendsto_infdist [tendsto_intros]:
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1857
  assumes f: "(f ---> l) F"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1858
  shows "((\<lambda>x. infdist (f x) A) ---> infdist l A) F"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1859
proof (rule tendstoI)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1860
  fix e ::real assume "0 < e"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1861
  from tendstoD[OF f this]
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1862
  show "eventually (\<lambda>x. dist (infdist (f x) A) (infdist l A) < e) F"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1863
  proof (eventually_elim)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1864
    fix x
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1865
    from infdist_triangle[of l A "f x"] infdist_triangle[of "f x" A l]
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1866
    have "dist (infdist (f x) A) (infdist l A) \<le> dist (f x) l"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1867
      by (simp add: dist_commute dist_real_def)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1868
    also assume "dist (f x) l < e"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1869
    finally show "dist (infdist (f x) A) (infdist l A) < e" .
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1870
  qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1871
qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1872
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1873
text{* Some other lemmas about sequences. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1874
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  1875
lemma sequentially_offset:
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  1876
  assumes "eventually (\<lambda>i. P i) sequentially"
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  1877
  shows "eventually (\<lambda>i. P (i + k)) sequentially"
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  1878
  using assms unfolding eventually_sequentially by (metis trans_le_add1)
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  1879
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1880
lemma seq_offset:
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  1881
  assumes "(f ---> l) sequentially"
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  1882
  shows "((\<lambda>i. f (i + k)) ---> l) sequentially"
44584
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  1883
  using assms by (rule LIMSEQ_ignore_initial_segment) (* FIXME: redundant *)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1884
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1885
lemma seq_offset_neg:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1886
  "(f ---> l) sequentially ==> ((\<lambda>i. f(i - k)) ---> l) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1887
  apply (rule topological_tendstoI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1888
  apply (drule (2) topological_tendstoD)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1889
  apply (simp only: eventually_sequentially)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1890
  apply (subgoal_tac "\<And>N k (n::nat). N + k <= n ==> N <= n - k")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1891
  apply metis
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1892
  by arith
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1893
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1894
lemma seq_offset_rev:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1895
  "((\<lambda>i. f(i + k)) ---> l) sequentially ==> (f ---> l) sequentially"
44584
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  1896
  by (rule LIMSEQ_offset) (* FIXME: redundant *)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1897
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1898
lemma seq_harmonic: "((\<lambda>n. inverse (real n)) ---> 0) sequentially"
44584
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  1899
  using LIMSEQ_inverse_real_of_nat by (rule LIMSEQ_imp_Suc)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1900
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1901
subsection {* More properties of closed balls *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1902
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1903
lemma closed_cball: "closed (cball x e)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1904
unfolding cball_def closed_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1905
unfolding Collect_neg_eq [symmetric] not_le
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1906
apply (clarsimp simp add: open_dist, rename_tac y)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1907
apply (rule_tac x="dist x y - e" in exI, clarsimp)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1908
apply (rename_tac x')
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1909
apply (cut_tac x=x and y=x' and z=y in dist_triangle)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1910
apply simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1911
done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1912
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1913
lemma open_contains_cball: "open S \<longleftrightarrow> (\<forall>x\<in>S. \<exists>e>0.  cball x e \<subseteq> S)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1914
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1915
  { fix x and e::real assume "x\<in>S" "e>0" "ball x e \<subseteq> S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1916
    hence "\<exists>d>0. cball x d \<subseteq> S" unfolding subset_eq by (rule_tac x="e/2" in exI, auto)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1917
  } moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1918
  { fix x and e::real assume "x\<in>S" "e>0" "cball x e \<subseteq> S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1919
    hence "\<exists>d>0. ball x d \<subseteq> S" unfolding subset_eq apply(rule_tac x="e/2" in exI) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1920
  } ultimately
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1921
  show ?thesis unfolding open_contains_ball by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1922
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1923
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1924
lemma open_contains_cball_eq: "open S ==> (\<forall>x. x \<in> S \<longleftrightarrow> (\<exists>e>0. cball x e \<subseteq> S))"
44170
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
  1925
  by (metis open_contains_cball subset_eq order_less_imp_le centre_in_cball)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1926
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1927
lemma mem_interior_cball: "x \<in> interior S \<longleftrightarrow> (\<exists>e>0. cball x e \<subseteq> S)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1928
  apply (simp add: interior_def, safe)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1929
  apply (force simp add: open_contains_cball)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1930
  apply (rule_tac x="ball x e" in exI)
36362
06475a1547cb fix lots of looping simp calls and other warnings
huffman
parents: 36360
diff changeset
  1931
  apply (simp add: subset_trans [OF ball_subset_cball])
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1932
  done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1933
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1934
lemma islimpt_ball:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1935
  fixes x y :: "'a::{real_normed_vector,perfect_space}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1936
  shows "y islimpt ball x e \<longleftrightarrow> 0 < e \<and> y \<in> cball x e" (is "?lhs = ?rhs")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1937
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1938
  assume "?lhs"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1939
  { assume "e \<le> 0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1940
    hence *:"ball x e = {}" using ball_eq_empty[of x e] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1941
    have False using `?lhs` unfolding * using islimpt_EMPTY[of y] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1942
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1943
  hence "e > 0" by (metis not_less)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1944
  moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1945
  have "y \<in> cball x e" using closed_cball[of x e] islimpt_subset[of y "ball x e" "cball x e"] ball_subset_cball[of x e] `?lhs` unfolding closed_limpt by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1946
  ultimately show "?rhs" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1947
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1948
  assume "?rhs" hence "e>0"  by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1949
  { fix d::real assume "d>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1950
    have "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1951
    proof(cases "d \<le> dist x y")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1952
      case True thus "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1953
      proof(cases "x=y")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1954
        case True hence False using `d \<le> dist x y` `d>0` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1955
        thus "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1956
      next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1957
        case False
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1958
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1959
        have "dist x (y - (d / (2 * dist y x)) *\<^sub>R (y - x))
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1960
              = norm (x - y + (d / (2 * norm (y - x))) *\<^sub>R (y - x))"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1961
          unfolding mem_cball mem_ball dist_norm diff_diff_eq2 diff_add_eq[THEN sym] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1962
        also have "\<dots> = \<bar>- 1 + d / (2 * norm (x - y))\<bar> * norm (x - y)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1963
          using scaleR_left_distrib[of "- 1" "d / (2 * norm (y - x))", THEN sym, of "y - x"]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1964
          unfolding scaleR_minus_left scaleR_one
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1965
          by (auto simp add: norm_minus_commute)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1966
        also have "\<dots> = \<bar>- norm (x - y) + d / 2\<bar>"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1967
          unfolding abs_mult_pos[of "norm (x - y)", OF norm_ge_zero[of "x - y"]]
49962
a8cc904a6820 Renamed {left,right}_distrib to distrib_{right,left}.
webertj
parents: 49834
diff changeset
  1968
          unfolding distrib_right using `x\<noteq>y`[unfolded dist_nz, unfolded dist_norm] by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1969
        also have "\<dots> \<le> e - d/2" using `d \<le> dist x y` and `d>0` and `?rhs` by(auto simp add: dist_norm)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1970
        finally have "y - (d / (2 * dist y x)) *\<^sub>R (y - x) \<in> ball x e" using `d>0` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1971
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1972
        moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1973
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1974
        have "(d / (2*dist y x)) *\<^sub>R (y - x) \<noteq> 0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1975
          using `x\<noteq>y`[unfolded dist_nz] `d>0` unfolding scaleR_eq_0_iff by (auto simp add: dist_commute)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1976
        moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1977
        have "dist (y - (d / (2 * dist y x)) *\<^sub>R (y - x)) y < d" unfolding dist_norm apply simp unfolding norm_minus_cancel
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1978
          using `d>0` `x\<noteq>y`[unfolded dist_nz] dist_commute[of x y]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1979
          unfolding dist_norm by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1980
        ultimately show "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d" by (rule_tac  x="y - (d / (2*dist y x)) *\<^sub>R (y - x)" in bexI) auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1981
      qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1982
    next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1983
      case False hence "d > dist x y" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1984
      show "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1985
      proof(cases "x=y")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1986
        case True
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1987
        obtain z where **: "z \<noteq> y" "dist z y < min e d"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1988
          using perfect_choose_dist[of "min e d" y]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1989
          using `d > 0` `e>0` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1990
        show "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1991
          unfolding `x = y`
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1992
          using `z \<noteq> y` **
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1993
          by (rule_tac x=z in bexI, auto simp add: dist_commute)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1994
      next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1995
        case False thus "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1996
          using `d>0` `d > dist x y` `?rhs` by(rule_tac x=x in bexI, auto)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1997
      qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1998
    qed  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1999
  thus "?lhs" unfolding mem_cball islimpt_approachable mem_ball by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2000
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2001
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2002
lemma closure_ball_lemma:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2003
  fixes x y :: "'a::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2004
  assumes "x \<noteq> y" shows "y islimpt ball x (dist x y)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2005
proof (rule islimptI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2006
  fix T assume "y \<in> T" "open T"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2007
  then obtain r where "0 < r" "\<forall>z. dist z y < r \<longrightarrow> z \<in> T"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2008
    unfolding open_dist by fast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2009
  (* choose point between x and y, within distance r of y. *)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2010
  def k \<equiv> "min 1 (r / (2 * dist x y))"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2011
  def z \<equiv> "y + scaleR k (x - y)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2012
  have z_def2: "z = x + scaleR (1 - k) (y - x)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2013
    unfolding z_def by (simp add: algebra_simps)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2014
  have "dist z y < r"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2015
    unfolding z_def k_def using `0 < r`
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2016
    by (simp add: dist_norm min_def)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2017
  hence "z \<in> T" using `\<forall>z. dist z y < r \<longrightarrow> z \<in> T` by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2018
  have "dist x z < dist x y"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2019
    unfolding z_def2 dist_norm
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2020
    apply (simp add: norm_minus_commute)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2021
    apply (simp only: dist_norm [symmetric])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2022
    apply (subgoal_tac "\<bar>1 - k\<bar> * dist x y < 1 * dist x y", simp)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2023
    apply (rule mult_strict_right_mono)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2024
    apply (simp add: k_def divide_pos_pos zero_less_dist_iff `0 < r` `x \<noteq> y`)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2025
    apply (simp add: zero_less_dist_iff `x \<noteq> y`)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2026
    done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2027
  hence "z \<in> ball x (dist x y)" by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2028
  have "z \<noteq> y"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2029
    unfolding z_def k_def using `x \<noteq> y` `0 < r`
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2030
    by (simp add: min_def)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2031
  show "\<exists>z\<in>ball x (dist x y). z \<in> T \<and> z \<noteq> y"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2032
    using `z \<in> ball x (dist x y)` `z \<in> T` `z \<noteq> y`
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2033
    by fast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2034
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2035
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2036
lemma closure_ball:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2037
  fixes x :: "'a::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2038
  shows "0 < e \<Longrightarrow> closure (ball x e) = cball x e"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2039
apply (rule equalityI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2040
apply (rule closure_minimal)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2041
apply (rule ball_subset_cball)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2042
apply (rule closed_cball)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2043
apply (rule subsetI, rename_tac y)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2044
apply (simp add: le_less [where 'a=real])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2045
apply (erule disjE)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2046
apply (rule subsetD [OF closure_subset], simp)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2047
apply (simp add: closure_def)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2048
apply clarify
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2049
apply (rule closure_ball_lemma)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2050
apply (simp add: zero_less_dist_iff)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2051
done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2052
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2053
(* In a trivial vector space, this fails for e = 0. *)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2054
lemma interior_cball:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2055
  fixes x :: "'a::{real_normed_vector, perfect_space}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2056
  shows "interior (cball x e) = ball x e"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2057
proof(cases "e\<ge>0")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2058
  case False note cs = this
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2059
  from cs have "ball x e = {}" using ball_empty[of e x] by auto moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2060
  { fix y assume "y \<in> cball x e"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2061
    hence False unfolding mem_cball using dist_nz[of x y] cs by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2062
  hence "cball x e = {}" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2063
  hence "interior (cball x e) = {}" using interior_empty by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2064
  ultimately show ?thesis by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2065
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2066
  case True note cs = this
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2067
  have "ball x e \<subseteq> cball x e" using ball_subset_cball by auto moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2068
  { fix S y assume as: "S \<subseteq> cball x e" "open S" "y\<in>S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2069
    then obtain d where "d>0" and d:"\<forall>x'. dist x' y < d \<longrightarrow> x' \<in> S" unfolding open_dist by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2070
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2071
    then obtain xa where xa_y: "xa \<noteq> y" and xa: "dist xa y < d"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2072
      using perfect_choose_dist [of d] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2073
    have "xa\<in>S" using d[THEN spec[where x=xa]] using xa by(auto simp add: dist_commute)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2074
    hence xa_cball:"xa \<in> cball x e" using as(1) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2075
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2076
    hence "y \<in> ball x e" proof(cases "x = y")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2077
      case True
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2078
      hence "e>0" using xa_y[unfolded dist_nz] xa_cball[unfolded mem_cball] by (auto simp add: dist_commute)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2079
      thus "y \<in> ball x e" using `x = y ` by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2080
    next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2081
      case False
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2082
      have "dist (y + (d / 2 / dist y x) *\<^sub>R (y - x)) y < d" unfolding dist_norm
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2083
        using `d>0` norm_ge_zero[of "y - x"] `x \<noteq> y` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2084
      hence *:"y + (d / 2 / dist y x) *\<^sub>R (y - x) \<in> cball x e" using d as(1)[unfolded subset_eq] by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2085
      have "y - x \<noteq> 0" using `x \<noteq> y` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2086
      hence **:"d / (2 * norm (y - x)) > 0" unfolding zero_less_norm_iff[THEN sym]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2087
        using `d>0` divide_pos_pos[of d "2*norm (y - x)"] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2088
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2089
      have "dist (y + (d / 2 / dist y x) *\<^sub>R (y - x)) x = norm (y + (d / (2 * norm (y - x))) *\<^sub>R y - (d / (2 * norm (y - x))) *\<^sub>R x - x)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2090
        by (auto simp add: dist_norm algebra_simps)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2091
      also have "\<dots> = norm ((1 + d / (2 * norm (y - x))) *\<^sub>R (y - x))"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2092
        by (auto simp add: algebra_simps)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2093
      also have "\<dots> = \<bar>1 + d / (2 * norm (y - x))\<bar> * norm (y - x)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2094
        using ** by auto
49962
a8cc904a6820 Renamed {left,right}_distrib to distrib_{right,left}.
webertj
parents: 49834
diff changeset
  2095
      also have "\<dots> = (dist y x) + d/2"using ** by (auto simp add: distrib_right dist_norm)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2096
      finally have "e \<ge> dist x y +d/2" using *[unfolded mem_cball] by (auto simp add: dist_commute)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2097
      thus "y \<in> ball x e" unfolding mem_ball using `d>0` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2098
    qed  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2099
  hence "\<forall>S \<subseteq> cball x e. open S \<longrightarrow> S \<subseteq> ball x e" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2100
  ultimately show ?thesis using interior_unique[of "ball x e" "cball x e"] using open_ball[of x e] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2101
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2102
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2103
lemma frontier_ball:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2104
  fixes a :: "'a::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2105
  shows "0 < e ==> frontier(ball a e) = {x. dist a x = e}"
36362
06475a1547cb fix lots of looping simp calls and other warnings
huffman
parents: 36360
diff changeset
  2106
  apply (simp add: frontier_def closure_ball interior_open order_less_imp_le)
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
  2107
  apply (simp add: set_eq_iff)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2108
  by arith
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2109
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2110
lemma frontier_cball:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2111
  fixes a :: "'a::{real_normed_vector, perfect_space}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2112
  shows "frontier(cball a e) = {x. dist a x = e}"
36362
06475a1547cb fix lots of looping simp calls and other warnings
huffman
parents: 36360
diff changeset
  2113
  apply (simp add: frontier_def interior_cball closed_cball order_less_imp_le)
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
  2114
  apply (simp add: set_eq_iff)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2115
  by arith
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2116
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2117
lemma cball_eq_empty: "(cball x e = {}) \<longleftrightarrow> e < 0"
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
  2118
  apply (simp add: set_eq_iff not_le)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2119
  by (metis zero_le_dist dist_self order_less_le_trans)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2120
lemma cball_empty: "e < 0 ==> cball x e = {}" by (simp add: cball_eq_empty)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2121
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2122
lemma cball_eq_sing:
44072
5b970711fb39 class perfect_space inherits from topological_space;
huffman
parents: 43338
diff changeset
  2123
  fixes x :: "'a::{metric_space,perfect_space}"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2124
  shows "(cball x e = {x}) \<longleftrightarrow> e = 0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2125
proof (rule linorder_cases)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2126
  assume e: "0 < e"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2127
  obtain a where "a \<noteq> x" "dist a x < e"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2128
    using perfect_choose_dist [OF e] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2129
  hence "a \<noteq> x" "dist x a \<le> e" by (auto simp add: dist_commute)
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
  2130
  with e show ?thesis by (auto simp add: set_eq_iff)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2131
qed auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2132
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2133
lemma cball_sing:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2134
  fixes x :: "'a::metric_space"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2135
  shows "e = 0 ==> cball x e = {x}"
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
  2136
  by (auto simp add: set_eq_iff)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2137
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  2138
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  2139
subsection {* Boundedness *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2140
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2141
  (* FIXME: This has to be unified with BSEQ!! *)
44207
ea99698c2070 locale-ize some definitions, so perfect_space and heine_borel can inherit from the proper superclasses
huffman
parents: 44170
diff changeset
  2142
definition (in metric_space)
ea99698c2070 locale-ize some definitions, so perfect_space and heine_borel can inherit from the proper superclasses
huffman
parents: 44170
diff changeset
  2143
  bounded :: "'a set \<Rightarrow> bool" where
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2144
  "bounded S \<longleftrightarrow> (\<exists>x e. \<forall>y\<in>S. dist x y \<le> e)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2145
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2146
lemma bounded_any_center: "bounded S \<longleftrightarrow> (\<exists>e. \<forall>y\<in>S. dist a y \<le> e)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2147
unfolding bounded_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2148
apply safe
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2149
apply (rule_tac x="dist a x + e" in exI, clarify)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2150
apply (drule (1) bspec)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2151
apply (erule order_trans [OF dist_triangle add_left_mono])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2152
apply auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2153
done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2154
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2155
lemma bounded_iff: "bounded S \<longleftrightarrow> (\<exists>a. \<forall>x\<in>S. norm x \<le> a)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2156
unfolding bounded_any_center [where a=0]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2157
by (simp add: dist_norm)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2158
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50094
diff changeset
  2159
lemma bounded_realI: assumes "\<forall>x\<in>s. abs (x::real) \<le> B" shows "bounded s"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50094
diff changeset
  2160
  unfolding bounded_def dist_real_def apply(rule_tac x=0 in exI)
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50094
diff changeset
  2161
  using assms by auto
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50094
diff changeset
  2162
50948
8c742f9de9f5 generalize topology lemmas; simplify proofs
huffman
parents: 50944
diff changeset
  2163
lemma bounded_empty [simp]: "bounded {}"
8c742f9de9f5 generalize topology lemmas; simplify proofs
huffman
parents: 50944
diff changeset
  2164
  by (simp add: bounded_def)
8c742f9de9f5 generalize topology lemmas; simplify proofs
huffman
parents: 50944
diff changeset
  2165
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2166
lemma bounded_subset: "bounded T \<Longrightarrow> S \<subseteq> T ==> bounded S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2167
  by (metis bounded_def subset_eq)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2168
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2169
lemma bounded_interior[intro]: "bounded S ==> bounded(interior S)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2170
  by (metis bounded_subset interior_subset)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2171
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2172
lemma bounded_closure[intro]: assumes "bounded S" shows "bounded(closure S)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2173
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2174
  from assms obtain x and a where a: "\<forall>y\<in>S. dist x y \<le> a" unfolding bounded_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2175
  { fix y assume "y \<in> closure S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2176
    then obtain f where f: "\<forall>n. f n \<in> S"  "(f ---> y) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2177
      unfolding closure_sequential by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2178
    have "\<forall>n. f n \<in> S \<longrightarrow> dist x (f n) \<le> a" using a by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2179
    hence "eventually (\<lambda>n. dist x (f n) \<le> a) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2180
      by (rule eventually_mono, simp add: f(1))
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2181
    have "dist x y \<le> a"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2182
      apply (rule Lim_dist_ubound [of sequentially f])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2183
      apply (rule trivial_limit_sequentially)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2184
      apply (rule f(2))
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2185
      apply fact
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2186
      done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2187
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2188
  thus ?thesis unfolding bounded_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2189
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2190
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2191
lemma bounded_cball[simp,intro]: "bounded (cball x e)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2192
  apply (simp add: bounded_def)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2193
  apply (rule_tac x=x in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2194
  apply (rule_tac x=e in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2195
  apply auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2196
  done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2197
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2198
lemma bounded_ball[simp,intro]: "bounded(ball x e)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2199
  by (metis ball_subset_cball bounded_cball bounded_subset)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2200
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2201
lemma bounded_Un[simp]: "bounded (S \<union> T) \<longleftrightarrow> bounded S \<and> bounded T"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2202
  apply (auto simp add: bounded_def)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2203
  apply (rename_tac x y r s)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2204
  apply (rule_tac x=x in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2205
  apply (rule_tac x="max r (dist x y + s)" in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2206
  apply (rule ballI, rename_tac z, safe)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2207
  apply (drule (1) bspec, simp)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2208
  apply (drule (1) bspec)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2209
  apply (rule min_max.le_supI2)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2210
  apply (erule order_trans [OF dist_triangle add_left_mono])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2211
  done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2212
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2213
lemma bounded_Union[intro]: "finite F \<Longrightarrow> (\<forall>S\<in>F. bounded S) \<Longrightarrow> bounded(\<Union>F)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2214
  by (induct rule: finite_induct[of F], auto)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2215
50955
ada575c605e1 simplify proof of compact_imp_bounded
huffman
parents: 50949
diff changeset
  2216
lemma bounded_UN [intro]: "finite A \<Longrightarrow> \<forall>x\<in>A. bounded (B x) \<Longrightarrow> bounded (\<Union>x\<in>A. B x)"
ada575c605e1 simplify proof of compact_imp_bounded
huffman
parents: 50949
diff changeset
  2217
  by (induct set: finite, auto)
ada575c605e1 simplify proof of compact_imp_bounded
huffman
parents: 50949
diff changeset
  2218
50948
8c742f9de9f5 generalize topology lemmas; simplify proofs
huffman
parents: 50944
diff changeset
  2219
lemma bounded_insert [simp]: "bounded (insert x S) \<longleftrightarrow> bounded S"
8c742f9de9f5 generalize topology lemmas; simplify proofs
huffman
parents: 50944
diff changeset
  2220
proof -
8c742f9de9f5 generalize topology lemmas; simplify proofs
huffman
parents: 50944
diff changeset
  2221
  have "\<forall>y\<in>{x}. dist x y \<le> 0" by simp
8c742f9de9f5 generalize topology lemmas; simplify proofs
huffman
parents: 50944
diff changeset
  2222
  hence "bounded {x}" unfolding bounded_def by fast
8c742f9de9f5 generalize topology lemmas; simplify proofs
huffman
parents: 50944
diff changeset
  2223
  thus ?thesis by (metis insert_is_Un bounded_Un)
8c742f9de9f5 generalize topology lemmas; simplify proofs
huffman
parents: 50944
diff changeset
  2224
qed
8c742f9de9f5 generalize topology lemmas; simplify proofs
huffman
parents: 50944
diff changeset
  2225
8c742f9de9f5 generalize topology lemmas; simplify proofs
huffman
parents: 50944
diff changeset
  2226
lemma finite_imp_bounded [intro]: "finite S \<Longrightarrow> bounded S"
8c742f9de9f5 generalize topology lemmas; simplify proofs
huffman
parents: 50944
diff changeset
  2227
  by (induct set: finite, simp_all)
8c742f9de9f5 generalize topology lemmas; simplify proofs
huffman
parents: 50944
diff changeset
  2228
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2229
lemma bounded_pos: "bounded S \<longleftrightarrow> (\<exists>b>0. \<forall>x\<in> S. norm x <= b)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2230
  apply (simp add: bounded_iff)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2231
  apply (subgoal_tac "\<And>x (y::real). 0 < 1 + abs y \<and> (x <= y \<longrightarrow> x <= 1 + abs y)")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2232
  by metis arith
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2233
50972
d2c6a0a7fcdf tuned proof
hoelzl
parents: 50971
diff changeset
  2234
lemma Bseq_eq_bounded: "Bseq f \<longleftrightarrow> bounded (range f)"
d2c6a0a7fcdf tuned proof
hoelzl
parents: 50971
diff changeset
  2235
  unfolding Bseq_def bounded_pos by auto
d2c6a0a7fcdf tuned proof
hoelzl
parents: 50971
diff changeset
  2236
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2237
lemma bounded_Int[intro]: "bounded S \<or> bounded T \<Longrightarrow> bounded (S \<inter> T)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2238
  by (metis Int_lower1 Int_lower2 bounded_subset)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2239
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2240
lemma bounded_diff[intro]: "bounded S ==> bounded (S - T)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2241
apply (metis Diff_subset bounded_subset)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2242
done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2243
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2244
lemma not_bounded_UNIV[simp, intro]:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2245
  "\<not> bounded (UNIV :: 'a::{real_normed_vector, perfect_space} set)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2246
proof(auto simp add: bounded_pos not_le)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2247
  obtain x :: 'a where "x \<noteq> 0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2248
    using perfect_choose_dist [OF zero_less_one] by fast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2249
  fix b::real  assume b: "b >0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2250
  have b1: "b +1 \<ge> 0" using b by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2251
  with `x \<noteq> 0` have "b < norm (scaleR (b + 1) (sgn x))"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2252
    by (simp add: norm_sgn)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2253
  then show "\<exists>x::'a. b < norm x" ..
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2254
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2255
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2256
lemma bounded_linear_image:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2257
  assumes "bounded S" "bounded_linear f"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2258
  shows "bounded(f ` S)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2259
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2260
  from assms(1) obtain b where b:"b>0" "\<forall>x\<in>S. norm x \<le> b" unfolding bounded_pos by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2261
  from assms(2) obtain B where B:"B>0" "\<forall>x. norm (f x) \<le> B * norm x" using bounded_linear.pos_bounded by (auto simp add: mult_ac)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2262
  { fix x assume "x\<in>S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2263
    hence "norm x \<le> b" using b by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2264
    hence "norm (f x) \<le> B * b" using B(2) apply(erule_tac x=x in allE)
36778
739a9379e29b avoid using real-specific versions of generic lemmas
huffman
parents: 36669
diff changeset
  2265
      by (metis B(1) B(2) order_trans mult_le_cancel_left_pos)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2266
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2267
  thus ?thesis unfolding bounded_pos apply(rule_tac x="b*B" in exI)
36778
739a9379e29b avoid using real-specific versions of generic lemmas
huffman
parents: 36669
diff changeset
  2268
    using b B mult_pos_pos [of b B] by (auto simp add: mult_commute)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2269
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2270
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2271
lemma bounded_scaling:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2272
  fixes S :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2273
  shows "bounded S \<Longrightarrow> bounded ((\<lambda>x. c *\<^sub>R x) ` S)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2274
  apply (rule bounded_linear_image, assumption)
44282
f0de18b62d63 remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents: 44252
diff changeset
  2275
  apply (rule bounded_linear_scaleR_right)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2276
  done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2277
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2278
lemma bounded_translation:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2279
  fixes S :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2280
  assumes "bounded S" shows "bounded ((\<lambda>x. a + x) ` S)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2281
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2282
  from assms obtain b where b:"b>0" "\<forall>x\<in>S. norm x \<le> b" unfolding bounded_pos by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2283
  { fix x assume "x\<in>S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2284
    hence "norm (a + x) \<le> b + norm a" using norm_triangle_ineq[of a x] b by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2285
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2286
  thus ?thesis unfolding bounded_pos using norm_ge_zero[of a] b(1) using add_strict_increasing[of b 0 "norm a"]
48048
87b94fb75198 remove stray reference to no-longer-existing theorem 'add'
huffman
parents: 47108
diff changeset
  2287
    by (auto intro!: exI[of _ "b + norm a"])
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2288
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2289
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2290
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2291
text{* Some theorems on sups and infs using the notion "bounded". *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2292
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2293
lemma bounded_real:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2294
  fixes S :: "real set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2295
  shows "bounded S \<longleftrightarrow>  (\<exists>a. \<forall>x\<in>S. abs x <= a)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2296
  by (simp add: bounded_iff)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2297
33270
paulson
parents: 33175
diff changeset
  2298
lemma bounded_has_Sup:
paulson
parents: 33175
diff changeset
  2299
  fixes S :: "real set"
paulson
parents: 33175
diff changeset
  2300
  assumes "bounded S" "S \<noteq> {}"
paulson
parents: 33175
diff changeset
  2301
  shows "\<forall>x\<in>S. x <= Sup S" and "\<forall>b. (\<forall>x\<in>S. x <= b) \<longrightarrow> Sup S <= b"
paulson
parents: 33175
diff changeset
  2302
proof
paulson
parents: 33175
diff changeset
  2303
  fix x assume "x\<in>S"
paulson
parents: 33175
diff changeset
  2304
  thus "x \<le> Sup S"
paulson
parents: 33175
diff changeset
  2305
    by (metis SupInf.Sup_upper abs_le_D1 assms(1) bounded_real)
paulson
parents: 33175
diff changeset
  2306
next
paulson
parents: 33175
diff changeset
  2307
  show "\<forall>b. (\<forall>x\<in>S. x \<le> b) \<longrightarrow> Sup S \<le> b" using assms
paulson
parents: 33175
diff changeset
  2308
    by (metis SupInf.Sup_least)
paulson
parents: 33175
diff changeset
  2309
qed
paulson
parents: 33175
diff changeset
  2310
paulson
parents: 33175
diff changeset
  2311
lemma Sup_insert:
paulson
parents: 33175
diff changeset
  2312
  fixes S :: "real set"
paulson
parents: 33175
diff changeset
  2313
  shows "bounded S ==> Sup(insert x S) = (if S = {} then x else max x (Sup S))" 
paulson
parents: 33175
diff changeset
  2314
by auto (metis Int_absorb Sup_insert_nonempty assms bounded_has_Sup(1) disjoint_iff_not_equal) 
paulson
parents: 33175
diff changeset
  2315
paulson
parents: 33175
diff changeset
  2316
lemma Sup_insert_finite:
paulson
parents: 33175
diff changeset
  2317
  fixes S :: "real set"
paulson
parents: 33175
diff changeset
  2318
  shows "finite S \<Longrightarrow> Sup(insert x S) = (if S = {} then x else max x (Sup S))"
paulson
parents: 33175
diff changeset
  2319
  apply (rule Sup_insert)
paulson
parents: 33175
diff changeset
  2320
  apply (rule finite_imp_bounded)
paulson
parents: 33175
diff changeset
  2321
  by simp
paulson
parents: 33175
diff changeset
  2322
paulson
parents: 33175
diff changeset
  2323
lemma bounded_has_Inf:
paulson
parents: 33175
diff changeset
  2324
  fixes S :: "real set"
paulson
parents: 33175
diff changeset
  2325
  assumes "bounded S"  "S \<noteq> {}"
paulson
parents: 33175
diff changeset
  2326
  shows "\<forall>x\<in>S. x >= Inf S" and "\<forall>b. (\<forall>x\<in>S. x >= b) \<longrightarrow> Inf S >= b"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2327
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2328
  fix x assume "x\<in>S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2329
  from assms(1) obtain a where a:"\<forall>x\<in>S. \<bar>x\<bar> \<le> a" unfolding bounded_real by auto
33270
paulson
parents: 33175
diff changeset
  2330
  thus "x \<ge> Inf S" using `x\<in>S`
paulson
parents: 33175
diff changeset
  2331
    by (metis Inf_lower_EX abs_le_D2 minus_le_iff)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2332
next
33270
paulson
parents: 33175
diff changeset
  2333
  show "\<forall>b. (\<forall>x\<in>S. x >= b) \<longrightarrow> Inf S \<ge> b" using assms
paulson
parents: 33175
diff changeset
  2334
    by (metis SupInf.Inf_greatest)
paulson
parents: 33175
diff changeset
  2335
qed
paulson
parents: 33175
diff changeset
  2336
paulson
parents: 33175
diff changeset
  2337
lemma Inf_insert:
paulson
parents: 33175
diff changeset
  2338
  fixes S :: "real set"
paulson
parents: 33175
diff changeset
  2339
  shows "bounded S ==> Inf(insert x S) = (if S = {} then x else min x (Inf S))" 
50944
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  2340
by auto (metis Int_absorb Inf_insert_nonempty bounded_has_Inf(1) disjoint_iff_not_equal)
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  2341
33270
paulson
parents: 33175
diff changeset
  2342
lemma Inf_insert_finite:
paulson
parents: 33175
diff changeset
  2343
  fixes S :: "real set"
paulson
parents: 33175
diff changeset
  2344
  shows "finite S ==> Inf(insert x S) = (if S = {} then x else min x (Inf S))"
paulson
parents: 33175
diff changeset
  2345
  by (rule Inf_insert, rule finite_imp_bounded, simp)
paulson
parents: 33175
diff changeset
  2346
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2347
subsection {* Compactness *}
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2348
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2349
subsubsection{* Open-cover compactness *}
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2350
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2351
definition compact :: "'a::topological_space set \<Rightarrow> bool" where
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2352
  compact_eq_heine_borel: -- "This name is used for backwards compatibility"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2353
    "compact S \<longleftrightarrow> (\<forall>C. (\<forall>c\<in>C. open c) \<and> S \<subseteq> \<Union>C \<longrightarrow> (\<exists>D\<subseteq>C. finite D \<and> S \<subseteq> \<Union>D))"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2354
50898
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2355
lemma compactI:
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2356
  assumes "\<And>C. \<forall>t\<in>C. open t \<Longrightarrow> s \<subseteq> \<Union> C \<Longrightarrow> \<exists>C'. C' \<subseteq> C \<and> finite C' \<and> s \<subseteq> \<Union> C'"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2357
  shows "compact s"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2358
  unfolding compact_eq_heine_borel using assms by metis
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2359
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2360
lemma compactE:
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2361
  assumes "compact s" and "\<forall>t\<in>C. open t" and "s \<subseteq> \<Union>C"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2362
  obtains C' where "C' \<subseteq> C" and "finite C'" and "s \<subseteq> \<Union>C'"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2363
  using assms unfolding compact_eq_heine_borel by metis
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2364
50944
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  2365
lemma compactE_image:
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  2366
  assumes "compact s" and "\<forall>t\<in>C. open (f t)" and "s \<subseteq> (\<Union>c\<in>C. f c)"
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  2367
  obtains C' where "C' \<subseteq> C" and "finite C'" and "s \<subseteq> (\<Union>c\<in>C'. f c)"
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  2368
  using assms unfolding ball_simps[symmetric] SUP_def
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  2369
  by (metis (lifting) finite_subset_image compact_eq_heine_borel[of s])
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  2370
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2371
subsubsection {* Bolzano-Weierstrass property *}
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2372
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2373
lemma heine_borel_imp_bolzano_weierstrass:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2374
  assumes "compact s" "infinite t"  "t \<subseteq> s"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2375
  shows "\<exists>x \<in> s. x islimpt t"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2376
proof(rule ccontr)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2377
  assume "\<not> (\<exists>x \<in> s. x islimpt t)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2378
  then obtain f where f:"\<forall>x\<in>s. x \<in> f x \<and> open (f x) \<and> (\<forall>y\<in>t. y \<in> f x \<longrightarrow> y = x)" unfolding islimpt_def
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2379
    using bchoice[of s "\<lambda> x T. x \<in> T \<and> open T \<and> (\<forall>y\<in>t. y \<in> T \<longrightarrow> y = x)"] by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2380
  obtain g where g:"g\<subseteq>{t. \<exists>x. x \<in> s \<and> t = f x}" "finite g" "s \<subseteq> \<Union>g"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2381
    using assms(1)[unfolded compact_eq_heine_borel, THEN spec[where x="{t. \<exists>x. x\<in>s \<and> t = f x}"]] using f by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2382
  from g(1,3) have g':"\<forall>x\<in>g. \<exists>xa \<in> s. x = f xa" by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2383
  { fix x y assume "x\<in>t" "y\<in>t" "f x = f y"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2384
    hence "x \<in> f x"  "y \<in> f x \<longrightarrow> y = x" using f[THEN bspec[where x=x]] and `t\<subseteq>s` by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2385
    hence "x = y" using `f x = f y` and f[THEN bspec[where x=y]] and `y\<in>t` and `t\<subseteq>s` by auto  }
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2386
  hence "inj_on f t" unfolding inj_on_def by simp
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2387
  hence "infinite (f ` t)" using assms(2) using finite_imageD by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2388
  moreover
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2389
  { fix x assume "x\<in>t" "f x \<notin> g"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2390
    from g(3) assms(3) `x\<in>t` obtain h where "h\<in>g" and "x\<in>h" by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2391
    then obtain y where "y\<in>s" "h = f y" using g'[THEN bspec[where x=h]] by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2392
    hence "y = x" using f[THEN bspec[where x=y]] and `x\<in>t` and `x\<in>h`[unfolded `h = f y`] by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2393
    hence False using `f x \<notin> g` `h\<in>g` unfolding `h = f y` by auto  }
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2394
  hence "f ` t \<subseteq> g" by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2395
  ultimately show False using g(2) using finite_subset by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2396
qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2397
50943
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  2398
lemma acc_point_range_imp_convergent_subsequence:
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  2399
  fixes l :: "'a :: first_countable_topology"
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  2400
  assumes l: "\<forall>U. l\<in>U \<longrightarrow> open U \<longrightarrow> infinite (U \<inter> range f)"
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  2401
  shows "\<exists>r. subseq r \<and> (f \<circ> r) ----> l"
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2402
proof -
50943
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  2403
  from countable_basis_at_decseq[of l] guess A . note A = this
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2404
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2405
  def s \<equiv> "\<lambda>n i. SOME j. i < j \<and> f j \<in> A (Suc n)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2406
  { fix n i
50943
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  2407
    have "infinite (A (Suc n) \<inter> range f - f`{.. i})"
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  2408
      using l A by auto
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  2409
    then have "\<exists>x. x \<in> A (Suc n) \<inter> range f - f`{.. i}"
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  2410
      unfolding ex_in_conv by (intro notI) simp
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  2411
    then have "\<exists>j. f j \<in> A (Suc n) \<and> j \<notin> {.. i}"
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  2412
      by auto
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  2413
    then have "\<exists>a. i < a \<and> f a \<in> A (Suc n)"
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  2414
      by (auto simp: not_le)
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2415
    then have "i < s n i" "f (s n i) \<in> A (Suc n)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2416
      unfolding s_def by (auto intro: someI2_ex) }
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2417
  note s = this
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2418
  def r \<equiv> "nat_rec (s 0 0) s"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2419
  have "subseq r"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2420
    by (auto simp: r_def s subseq_Suc_iff)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2421
  moreover
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2422
  have "(\<lambda>n. f (r n)) ----> l"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2423
  proof (rule topological_tendstoI)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2424
    fix S assume "open S" "l \<in> S"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2425
    with A(3) have "eventually (\<lambda>i. A i \<subseteq> S) sequentially" by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2426
    moreover
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2427
    { fix i assume "Suc 0 \<le> i" then have "f (r i) \<in> A i"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2428
        by (cases i) (simp_all add: r_def s) }
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2429
    then have "eventually (\<lambda>i. f (r i) \<in> A i) sequentially" by (auto simp: eventually_sequentially)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2430
    ultimately show "eventually (\<lambda>i. f (r i) \<in> S) sequentially"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2431
      by eventually_elim auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2432
  qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2433
  ultimately show "\<exists>r. subseq r \<and> (f \<circ> r) ----> l"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2434
    by (auto simp: convergent_def comp_def)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2435
qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2436
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2437
lemma sequence_infinite_lemma:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2438
  fixes f :: "nat \<Rightarrow> 'a::t1_space"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2439
  assumes "\<forall>n. f n \<noteq> l" and "(f ---> l) sequentially"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2440
  shows "infinite (range f)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2441
proof
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2442
  assume "finite (range f)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2443
  hence "closed (range f)" by (rule finite_imp_closed)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2444
  hence "open (- range f)" by (rule open_Compl)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2445
  from assms(1) have "l \<in> - range f" by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2446
  from assms(2) have "eventually (\<lambda>n. f n \<in> - range f) sequentially"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2447
    using `open (- range f)` `l \<in> - range f` by (rule topological_tendstoD)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2448
  thus False unfolding eventually_sequentially by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2449
qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2450
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2451
lemma closure_insert:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2452
  fixes x :: "'a::t1_space"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2453
  shows "closure (insert x s) = insert x (closure s)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2454
apply (rule closure_unique)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2455
apply (rule insert_mono [OF closure_subset])
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2456
apply (rule closed_insert [OF closed_closure])
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2457
apply (simp add: closure_minimal)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2458
done
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2459
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2460
lemma islimpt_insert:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2461
  fixes x :: "'a::t1_space"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2462
  shows "x islimpt (insert a s) \<longleftrightarrow> x islimpt s"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2463
proof
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2464
  assume *: "x islimpt (insert a s)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2465
  show "x islimpt s"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2466
  proof (rule islimptI)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2467
    fix t assume t: "x \<in> t" "open t"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2468
    show "\<exists>y\<in>s. y \<in> t \<and> y \<noteq> x"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2469
    proof (cases "x = a")
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2470
      case True
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2471
      obtain y where "y \<in> insert a s" "y \<in> t" "y \<noteq> x"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2472
        using * t by (rule islimptE)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2473
      with `x = a` show ?thesis by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2474
    next
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2475
      case False
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2476
      with t have t': "x \<in> t - {a}" "open (t - {a})"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2477
        by (simp_all add: open_Diff)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2478
      obtain y where "y \<in> insert a s" "y \<in> t - {a}" "y \<noteq> x"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2479
        using * t' by (rule islimptE)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2480
      thus ?thesis by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2481
    qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2482
  qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2483
next
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2484
  assume "x islimpt s" thus "x islimpt (insert a s)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2485
    by (rule islimpt_subset) auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2486
qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2487
50897
078590669527 generalize lemma islimpt_finite to class t1_space
huffman
parents: 50884
diff changeset
  2488
lemma islimpt_finite:
078590669527 generalize lemma islimpt_finite to class t1_space
huffman
parents: 50884
diff changeset
  2489
  fixes x :: "'a::t1_space"
078590669527 generalize lemma islimpt_finite to class t1_space
huffman
parents: 50884
diff changeset
  2490
  shows "finite s \<Longrightarrow> \<not> x islimpt s"
078590669527 generalize lemma islimpt_finite to class t1_space
huffman
parents: 50884
diff changeset
  2491
by (induct set: finite, simp_all add: islimpt_insert)
078590669527 generalize lemma islimpt_finite to class t1_space
huffman
parents: 50884
diff changeset
  2492
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2493
lemma islimpt_union_finite:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2494
  fixes x :: "'a::t1_space"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2495
  shows "finite s \<Longrightarrow> x islimpt (s \<union> t) \<longleftrightarrow> x islimpt t"
50897
078590669527 generalize lemma islimpt_finite to class t1_space
huffman
parents: 50884
diff changeset
  2496
by (simp add: islimpt_Un islimpt_finite)
078590669527 generalize lemma islimpt_finite to class t1_space
huffman
parents: 50884
diff changeset
  2497
50943
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  2498
lemma islimpt_eq_acc_point:
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  2499
  fixes l :: "'a :: t1_space"
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  2500
  shows "l islimpt S \<longleftrightarrow> (\<forall>U. l\<in>U \<longrightarrow> open U \<longrightarrow> infinite (U \<inter> S))"
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  2501
proof (safe intro!: islimptI)
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  2502
  fix U assume "l islimpt S" "l \<in> U" "open U" "finite (U \<inter> S)"
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  2503
  then have "l islimpt S" "l \<in> (U - (U \<inter> S - {l}))" "open (U - (U \<inter> S - {l}))"
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  2504
    by (auto intro: finite_imp_closed)
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  2505
  then show False
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  2506
    by (rule islimptE) auto
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  2507
next
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  2508
  fix T assume *: "\<forall>U. l\<in>U \<longrightarrow> open U \<longrightarrow> infinite (U \<inter> S)" "l \<in> T" "open T"
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  2509
  then have "infinite (T \<inter> S - {l})" by auto
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  2510
  then have "\<exists>x. x \<in> (T \<inter> S - {l})"
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  2511
    unfolding ex_in_conv by (intro notI) simp
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  2512
  then show "\<exists>y\<in>S. y \<in> T \<and> y \<noteq> l"
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  2513
    by auto
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  2514
qed
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  2515
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  2516
lemma islimpt_range_imp_convergent_subsequence:
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  2517
  fixes l :: "'a :: {t1_space, first_countable_topology}"
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  2518
  assumes l: "l islimpt (range f)"
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  2519
  shows "\<exists>r. subseq r \<and> (f \<circ> r) ----> l"
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  2520
  using l unfolding islimpt_eq_acc_point
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  2521
  by (rule acc_point_range_imp_convergent_subsequence)
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  2522
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2523
lemma sequence_unique_limpt:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2524
  fixes f :: "nat \<Rightarrow> 'a::t2_space"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2525
  assumes "(f ---> l) sequentially" and "l' islimpt (range f)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2526
  shows "l' = l"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2527
proof (rule ccontr)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2528
  assume "l' \<noteq> l"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2529
  obtain s t where "open s" "open t" "l' \<in> s" "l \<in> t" "s \<inter> t = {}"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2530
    using hausdorff [OF `l' \<noteq> l`] by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2531
  have "eventually (\<lambda>n. f n \<in> t) sequentially"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2532
    using assms(1) `open t` `l \<in> t` by (rule topological_tendstoD)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2533
  then obtain N where "\<forall>n\<ge>N. f n \<in> t"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2534
    unfolding eventually_sequentially by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2535
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2536
  have "UNIV = {..<N} \<union> {N..}" by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2537
  hence "l' islimpt (f ` ({..<N} \<union> {N..}))" using assms(2) by simp
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2538
  hence "l' islimpt (f ` {..<N} \<union> f ` {N..})" by (simp add: image_Un)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2539
  hence "l' islimpt (f ` {N..})" by (simp add: islimpt_union_finite)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2540
  then obtain y where "y \<in> f ` {N..}" "y \<in> s" "y \<noteq> l'"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2541
    using `l' \<in> s` `open s` by (rule islimptE)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2542
  then obtain n where "N \<le> n" "f n \<in> s" "f n \<noteq> l'" by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2543
  with `\<forall>n\<ge>N. f n \<in> t` have "f n \<in> s \<inter> t" by simp
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2544
  with `s \<inter> t = {}` show False by simp
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2545
qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2546
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2547
lemma bolzano_weierstrass_imp_closed:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2548
  fixes s :: "'a::{first_countable_topology, t2_space} set"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2549
  assumes "\<forall>t. infinite t \<and> t \<subseteq> s --> (\<exists>x \<in> s. x islimpt t)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2550
  shows "closed s"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2551
proof-
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2552
  { fix x l assume as: "\<forall>n::nat. x n \<in> s" "(x ---> l) sequentially"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2553
    hence "l \<in> s"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2554
    proof(cases "\<forall>n. x n \<noteq> l")
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2555
      case False thus "l\<in>s" using as(1) by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2556
    next
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2557
      case True note cas = this
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2558
      with as(2) have "infinite (range x)" using sequence_infinite_lemma[of x l] by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2559
      then obtain l' where "l'\<in>s" "l' islimpt (range x)" using assms[THEN spec[where x="range x"]] as(1) by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2560
      thus "l\<in>s" using sequence_unique_limpt[of x l l'] using as cas by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2561
    qed  }
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2562
  thus ?thesis unfolding closed_sequential_limits by fast
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2563
qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2564
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2565
lemma compact_imp_closed:
50898
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2566
  fixes s :: "'a::t2_space set"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2567
  assumes "compact s" shows "closed s"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2568
unfolding closed_def
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2569
proof (rule openI)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2570
  fix y assume "y \<in> - s"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2571
  let ?C = "\<Union>x\<in>s. {u. open u \<and> x \<in> u \<and> eventually (\<lambda>y. y \<notin> u) (nhds y)}"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2572
  note `compact s`
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2573
  moreover have "\<forall>u\<in>?C. open u" by simp
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2574
  moreover have "s \<subseteq> \<Union>?C"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2575
  proof
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2576
    fix x assume "x \<in> s"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2577
    with `y \<in> - s` have "x \<noteq> y" by clarsimp
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2578
    hence "\<exists>u v. open u \<and> open v \<and> x \<in> u \<and> y \<in> v \<and> u \<inter> v = {}"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2579
      by (rule hausdorff)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2580
    with `x \<in> s` show "x \<in> \<Union>?C"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2581
      unfolding eventually_nhds by auto
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2582
  qed
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2583
  ultimately obtain D where "D \<subseteq> ?C" and "finite D" and "s \<subseteq> \<Union>D"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2584
    by (rule compactE)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2585
  from `D \<subseteq> ?C` have "\<forall>x\<in>D. eventually (\<lambda>y. y \<notin> x) (nhds y)" by auto
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2586
  with `finite D` have "eventually (\<lambda>y. y \<notin> \<Union>D) (nhds y)"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2587
    by (simp add: eventually_Ball_finite)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2588
  with `s \<subseteq> \<Union>D` have "eventually (\<lambda>y. y \<notin> s) (nhds y)"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2589
    by (auto elim!: eventually_mono [rotated])
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2590
  thus "\<exists>t. open t \<and> y \<in> t \<and> t \<subseteq> - s"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2591
    by (simp add: eventually_nhds subset_eq)
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2592
qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2593
50944
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  2594
lemma compact_imp_bounded:
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  2595
  assumes "compact U" shows "bounded U"
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  2596
proof -
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  2597
  have "compact U" "\<forall>x\<in>U. open (ball x 1)" "U \<subseteq> (\<Union>x\<in>U. ball x 1)" using assms by auto
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  2598
  then obtain D where D: "D \<subseteq> U" "finite D" "U \<subseteq> (\<Union>x\<in>D. ball x 1)"
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  2599
    by (elim compactE_image)
50955
ada575c605e1 simplify proof of compact_imp_bounded
huffman
parents: 50949
diff changeset
  2600
  from `finite D` have "bounded (\<Union>x\<in>D. ball x 1)"
ada575c605e1 simplify proof of compact_imp_bounded
huffman
parents: 50949
diff changeset
  2601
    by (simp add: bounded_UN)
ada575c605e1 simplify proof of compact_imp_bounded
huffman
parents: 50949
diff changeset
  2602
  thus "bounded U" using `U \<subseteq> (\<Union>x\<in>D. ball x 1)` 
ada575c605e1 simplify proof of compact_imp_bounded
huffman
parents: 50949
diff changeset
  2603
    by (rule bounded_subset)
50944
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  2604
qed
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  2605
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2606
text{* In particular, some common special cases. *}
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2607
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2608
lemma compact_empty[simp]:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2609
 "compact {}"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2610
  unfolding compact_eq_heine_borel
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2611
  by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2612
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2613
lemma compact_union [intro]:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2614
  assumes "compact s" "compact t" shows " compact (s \<union> t)"
50898
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2615
proof (rule compactI)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2616
  fix f assume *: "Ball f open" "s \<union> t \<subseteq> \<Union>f"
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2617
  from * `compact s` obtain s' where "s' \<subseteq> f \<and> finite s' \<and> s \<subseteq> \<Union>s'"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2618
    unfolding compact_eq_heine_borel by (auto elim!: allE[of _ f]) metis
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2619
  moreover from * `compact t` obtain t' where "t' \<subseteq> f \<and> finite t' \<and> t \<subseteq> \<Union>t'"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2620
    unfolding compact_eq_heine_borel by (auto elim!: allE[of _ f]) metis
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2621
  ultimately show "\<exists>f'\<subseteq>f. finite f' \<and> s \<union> t \<subseteq> \<Union>f'"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2622
    by (auto intro!: exI[of _ "s' \<union> t'"])
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2623
qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2624
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2625
lemma compact_Union [intro]: "finite S \<Longrightarrow> (\<And>T. T \<in> S \<Longrightarrow> compact T) \<Longrightarrow> compact (\<Union>S)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2626
  by (induct set: finite) auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2627
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2628
lemma compact_UN [intro]:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2629
  "finite A \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> compact (B x)) \<Longrightarrow> compact (\<Union>x\<in>A. B x)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2630
  unfolding SUP_def by (rule compact_Union) auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2631
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2632
lemma compact_inter_closed [intro]:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2633
  assumes "compact s" and "closed t"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2634
  shows "compact (s \<inter> t)"
50898
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2635
proof (rule compactI)
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2636
  fix C assume C: "\<forall>c\<in>C. open c" and cover: "s \<inter> t \<subseteq> \<Union>C"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2637
  from C `closed t` have "\<forall>c\<in>C \<union> {-t}. open c" by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2638
  moreover from cover have "s \<subseteq> \<Union>(C \<union> {-t})" by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2639
  ultimately have "\<exists>D\<subseteq>C \<union> {-t}. finite D \<and> s \<subseteq> \<Union>D"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2640
    using `compact s` unfolding compact_eq_heine_borel by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2641
  then guess D ..
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2642
  then show "\<exists>D\<subseteq>C. finite D \<and> s \<inter> t \<subseteq> \<Union>D"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2643
    by (intro exI[of _ "D - {-t}"]) auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2644
qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2645
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2646
lemma closed_inter_compact [intro]:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2647
  assumes "closed s" and "compact t"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2648
  shows "compact (s \<inter> t)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2649
  using compact_inter_closed [of t s] assms
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2650
  by (simp add: Int_commute)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2651
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2652
lemma compact_inter [intro]:
50898
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2653
  fixes s t :: "'a :: t2_space set"
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2654
  assumes "compact s" and "compact t"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2655
  shows "compact (s \<inter> t)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2656
  using assms by (intro compact_inter_closed compact_imp_closed)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2657
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2658
lemma compact_sing [simp]: "compact {a}"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2659
  unfolding compact_eq_heine_borel by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2660
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2661
lemma compact_insert [simp]:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2662
  assumes "compact s" shows "compact (insert x s)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2663
proof -
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2664
  have "compact ({x} \<union> s)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2665
    using compact_sing assms by (rule compact_union)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2666
  thus ?thesis by simp
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2667
qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2668
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2669
lemma finite_imp_compact:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2670
  shows "finite s \<Longrightarrow> compact s"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2671
  by (induct set: finite) simp_all
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2672
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2673
lemma open_delete:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2674
  fixes s :: "'a::t1_space set"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2675
  shows "open s \<Longrightarrow> open (s - {x})"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2676
  by (simp add: open_Diff)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2677
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2678
text{* Finite intersection property *}
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2679
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2680
lemma inj_setminus: "inj_on uminus (A::'a set set)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2681
  by (auto simp: inj_on_def)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2682
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2683
lemma compact_fip:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2684
  "compact U \<longleftrightarrow>
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2685
    (\<forall>A. (\<forall>a\<in>A. closed a) \<longrightarrow> (\<forall>B \<subseteq> A. finite B \<longrightarrow> U \<inter> \<Inter>B \<noteq> {}) \<longrightarrow> U \<inter> \<Inter>A \<noteq> {})"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2686
  (is "_ \<longleftrightarrow> ?R")
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2687
proof (safe intro!: compact_eq_heine_borel[THEN iffD2])
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2688
  fix A assume "compact U" and A: "\<forall>a\<in>A. closed a" "U \<inter> \<Inter>A = {}"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2689
    and fi: "\<forall>B \<subseteq> A. finite B \<longrightarrow> U \<inter> \<Inter>B \<noteq> {}"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2690
  from A have "(\<forall>a\<in>uminus`A. open a) \<and> U \<subseteq> \<Union>uminus`A"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2691
    by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2692
  with `compact U` obtain B where "B \<subseteq> A" "finite (uminus`B)" "U \<subseteq> \<Union>(uminus`B)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2693
    unfolding compact_eq_heine_borel by (metis subset_image_iff)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2694
  with fi[THEN spec, of B] show False
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2695
    by (auto dest: finite_imageD intro: inj_setminus)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2696
next
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2697
  fix A assume ?R and cover: "\<forall>a\<in>A. open a" "U \<subseteq> \<Union>A"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2698
  from cover have "U \<inter> \<Inter>(uminus`A) = {}" "\<forall>a\<in>uminus`A. closed a"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2699
    by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2700
  with `?R` obtain B where "B \<subseteq> A" "finite (uminus`B)" "U \<inter> \<Inter>uminus`B = {}"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2701
    by (metis subset_image_iff)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2702
  then show "\<exists>T\<subseteq>A. finite T \<and> U \<subseteq> \<Union>T"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2703
    by  (auto intro!: exI[of _ B] inj_setminus dest: finite_imageD)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2704
qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2705
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2706
lemma compact_imp_fip:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2707
  "compact s \<Longrightarrow> \<forall>t \<in> f. closed t \<Longrightarrow> \<forall>f'. finite f' \<and> f' \<subseteq> f \<longrightarrow> (s \<inter> (\<Inter> f') \<noteq> {}) \<Longrightarrow>
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2708
    s \<inter> (\<Inter> f) \<noteq> {}"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2709
  unfolding compact_fip by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2710
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2711
text{*Compactness expressed with filters*}
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2712
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2713
definition "filter_from_subbase B = Abs_filter (\<lambda>P. \<exists>X \<subseteq> B. finite X \<and> Inf X \<le> P)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2714
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2715
lemma eventually_filter_from_subbase:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2716
  "eventually P (filter_from_subbase B) \<longleftrightarrow> (\<exists>X \<subseteq> B. finite X \<and> Inf X \<le> P)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2717
    (is "_ \<longleftrightarrow> ?R P")
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2718
  unfolding filter_from_subbase_def
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2719
proof (rule eventually_Abs_filter is_filter.intro)+
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2720
  show "?R (\<lambda>x. True)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2721
    by (rule exI[of _ "{}"]) (simp add: le_fun_def)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2722
next
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2723
  fix P Q assume "?R P" then guess X ..
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2724
  moreover assume "?R Q" then guess Y ..
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2725
  ultimately show "?R (\<lambda>x. P x \<and> Q x)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2726
    by (intro exI[of _ "X \<union> Y"]) auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2727
next
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2728
  fix P Q
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2729
  assume "?R P" then guess X ..
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2730
  moreover assume "\<forall>x. P x \<longrightarrow> Q x"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2731
  ultimately show "?R Q"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2732
    by (intro exI[of _ X]) auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2733
qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2734
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2735
lemma eventually_filter_from_subbaseI: "P \<in> B \<Longrightarrow> eventually P (filter_from_subbase B)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2736
  by (subst eventually_filter_from_subbase) (auto intro!: exI[of _ "{P}"])
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2737
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2738
lemma filter_from_subbase_not_bot:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2739
  "\<forall>X \<subseteq> B. finite X \<longrightarrow> Inf X \<noteq> bot \<Longrightarrow> filter_from_subbase B \<noteq> bot"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2740
  unfolding trivial_limit_def eventually_filter_from_subbase by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2741
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2742
lemma closure_iff_nhds_not_empty:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2743
  "x \<in> closure X \<longleftrightarrow> (\<forall>A. \<forall>S\<subseteq>A. open S \<longrightarrow> x \<in> S \<longrightarrow> X \<inter> A \<noteq> {})"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2744
proof safe
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2745
  assume x: "x \<in> closure X"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2746
  fix S A assume "open S" "x \<in> S" "X \<inter> A = {}" "S \<subseteq> A"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2747
  then have "x \<notin> closure (-S)" 
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2748
    by (auto simp: closure_complement subset_eq[symmetric] intro: interiorI)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2749
  with x have "x \<in> closure X - closure (-S)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2750
    by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2751
  also have "\<dots> \<subseteq> closure (X \<inter> S)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2752
    using `open S` open_inter_closure_subset[of S X] by (simp add: closed_Compl ac_simps)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2753
  finally have "X \<inter> S \<noteq> {}" by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2754
  then show False using `X \<inter> A = {}` `S \<subseteq> A` by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2755
next
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2756
  assume "\<forall>A S. S \<subseteq> A \<longrightarrow> open S \<longrightarrow> x \<in> S \<longrightarrow> X \<inter> A \<noteq> {}"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2757
  from this[THEN spec, of "- X", THEN spec, of "- closure X"]
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2758
  show "x \<in> closure X"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2759
    by (simp add: closure_subset open_Compl)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2760
qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2761
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2762
lemma compact_filter:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2763
  "compact U \<longleftrightarrow> (\<forall>F. F \<noteq> bot \<longrightarrow> eventually (\<lambda>x. x \<in> U) F \<longrightarrow> (\<exists>x\<in>U. inf (nhds x) F \<noteq> bot))"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2764
proof (intro allI iffI impI compact_fip[THEN iffD2] notI)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2765
  fix F assume "compact U" and F: "F \<noteq> bot" "eventually (\<lambda>x. x \<in> U) F"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2766
  from F have "U \<noteq> {}"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2767
    by (auto simp: eventually_False)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2768
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2769
  def Z \<equiv> "closure ` {A. eventually (\<lambda>x. x \<in> A) F}"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2770
  then have "\<forall>z\<in>Z. closed z"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2771
    by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2772
  moreover 
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2773
  have ev_Z: "\<And>z. z \<in> Z \<Longrightarrow> eventually (\<lambda>x. x \<in> z) F"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2774
    unfolding Z_def by (auto elim: eventually_elim1 intro: set_mp[OF closure_subset])
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2775
  have "(\<forall>B \<subseteq> Z. finite B \<longrightarrow> U \<inter> \<Inter>B \<noteq> {})"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2776
  proof (intro allI impI)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2777
    fix B assume "finite B" "B \<subseteq> Z"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2778
    with `finite B` ev_Z have "eventually (\<lambda>x. \<forall>b\<in>B. x \<in> b) F"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2779
      by (auto intro!: eventually_Ball_finite)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2780
    with F(2) have "eventually (\<lambda>x. x \<in> U \<inter> (\<Inter>B)) F"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2781
      by eventually_elim auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2782
    with F show "U \<inter> \<Inter>B \<noteq> {}"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2783
      by (intro notI) (simp add: eventually_False)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2784
  qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2785
  ultimately have "U \<inter> \<Inter>Z \<noteq> {}"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2786
    using `compact U` unfolding compact_fip by blast
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2787
  then obtain x where "x \<in> U" and x: "\<And>z. z \<in> Z \<Longrightarrow> x \<in> z" by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2788
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2789
  have "\<And>P. eventually P (inf (nhds x) F) \<Longrightarrow> P \<noteq> bot"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2790
    unfolding eventually_inf eventually_nhds
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2791
  proof safe
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2792
    fix P Q R S
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2793
    assume "eventually R F" "open S" "x \<in> S"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2794
    with open_inter_closure_eq_empty[of S "{x. R x}"] x[of "closure {x. R x}"]
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2795
    have "S \<inter> {x. R x} \<noteq> {}" by (auto simp: Z_def)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2796
    moreover assume "Ball S Q" "\<forall>x. Q x \<and> R x \<longrightarrow> bot x"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2797
    ultimately show False by (auto simp: set_eq_iff)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2798
  qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2799
  with `x \<in> U` show "\<exists>x\<in>U. inf (nhds x) F \<noteq> bot"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2800
    by (metis eventually_bot)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2801
next
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2802
  fix A assume A: "\<forall>a\<in>A. closed a" "\<forall>B\<subseteq>A. finite B \<longrightarrow> U \<inter> \<Inter>B \<noteq> {}" "U \<inter> \<Inter>A = {}"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2803
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2804
  def P' \<equiv> "(\<lambda>a (x::'a). x \<in> a)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2805
  then have inj_P': "\<And>A. inj_on P' A"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2806
    by (auto intro!: inj_onI simp: fun_eq_iff)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2807
  def F \<equiv> "filter_from_subbase (P' ` insert U A)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2808
  have "F \<noteq> bot"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2809
    unfolding F_def
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2810
  proof (safe intro!: filter_from_subbase_not_bot)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2811
    fix X assume "X \<subseteq> P' ` insert U A" "finite X" "Inf X = bot"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2812
    then obtain B where "B \<subseteq> insert U A" "finite B" and B: "Inf (P' ` B) = bot"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2813
      unfolding subset_image_iff by (auto intro: inj_P' finite_imageD)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2814
    with A(2)[THEN spec, of "B - {U}"] have "U \<inter> \<Inter>(B - {U}) \<noteq> {}" by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2815
    with B show False by (auto simp: P'_def fun_eq_iff)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2816
  qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2817
  moreover have "eventually (\<lambda>x. x \<in> U) F"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2818
    unfolding F_def by (rule eventually_filter_from_subbaseI) (auto simp: P'_def)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2819
  moreover assume "\<forall>F. F \<noteq> bot \<longrightarrow> eventually (\<lambda>x. x \<in> U) F \<longrightarrow> (\<exists>x\<in>U. inf (nhds x) F \<noteq> bot)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2820
  ultimately obtain x where "x \<in> U" and x: "inf (nhds x) F \<noteq> bot"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2821
    by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2822
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2823
  { fix V assume "V \<in> A"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2824
    then have V: "eventually (\<lambda>x. x \<in> V) F"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2825
      by (auto simp add: F_def image_iff P'_def intro!: eventually_filter_from_subbaseI)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2826
    have "x \<in> closure V"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2827
      unfolding closure_iff_nhds_not_empty
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2828
    proof (intro impI allI)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2829
      fix S A assume "open S" "x \<in> S" "S \<subseteq> A"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2830
      then have "eventually (\<lambda>x. x \<in> A) (nhds x)" by (auto simp: eventually_nhds)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2831
      with V have "eventually (\<lambda>x. x \<in> V \<inter> A) (inf (nhds x) F)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2832
        by (auto simp: eventually_inf)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2833
      with x show "V \<inter> A \<noteq> {}"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2834
        by (auto simp del: Int_iff simp add: trivial_limit_def)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2835
    qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2836
    then have "x \<in> V"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2837
      using `V \<in> A` A(1) by simp }
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2838
  with `x\<in>U` have "x \<in> U \<inter> \<Inter>A" by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2839
  with `U \<inter> \<Inter>A = {}` show False by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2840
qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2841
50941
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2842
definition "countably_compact U \<longleftrightarrow>
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2843
    (\<forall>A. countable A \<longrightarrow> (\<forall>a\<in>A. open a) \<longrightarrow> U \<subseteq> \<Union>A \<longrightarrow> (\<exists>T\<subseteq>A. finite T \<and> U \<subseteq> \<Union>T))"
50941
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2844
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2845
lemma countably_compactE:
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2846
  assumes "countably_compact s" and "\<forall>t\<in>C. open t" and "s \<subseteq> \<Union>C" "countable C"
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2847
  obtains C' where "C' \<subseteq> C" and "finite C'" and "s \<subseteq> \<Union>C'"
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2848
  using assms unfolding countably_compact_def by metis
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2849
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2850
lemma countably_compactI:
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2851
  assumes "\<And>C. \<forall>t\<in>C. open t \<Longrightarrow> s \<subseteq> \<Union>C \<Longrightarrow> countable C \<Longrightarrow> (\<exists>C'\<subseteq>C. finite C' \<and> s \<subseteq> \<Union>C')"
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2852
  shows "countably_compact s"
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2853
  using assms unfolding countably_compact_def by metis
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2854
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2855
lemma compact_imp_countably_compact: "compact U \<Longrightarrow> countably_compact U"
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2856
  by (auto simp: compact_eq_heine_borel countably_compact_def)
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2857
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2858
lemma countably_compact_imp_compact:
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2859
  assumes "countably_compact U"
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2860
  assumes ccover: "countable B" "\<forall>b\<in>B. open b"
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2861
  assumes basis: "\<And>T x. open T \<Longrightarrow> x \<in> T \<Longrightarrow> x \<in> U \<Longrightarrow> \<exists>b\<in>B. x \<in> b \<and> b \<inter> U \<subseteq> T"
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2862
  shows "compact U"
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2863
  using `countably_compact U` unfolding compact_eq_heine_borel countably_compact_def
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2864
proof safe
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2865
  fix A assume A: "\<forall>a\<in>A. open a" "U \<subseteq> \<Union>A"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2866
  assume *: "\<forall>A. countable A \<longrightarrow> (\<forall>a\<in>A. open a) \<longrightarrow> U \<subseteq> \<Union>A \<longrightarrow> (\<exists>T\<subseteq>A. finite T \<and> U \<subseteq> \<Union>T)"
50941
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2867
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2868
  moreover def C \<equiv> "{b\<in>B. \<exists>a\<in>A. b \<inter> U \<subseteq> a}"
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2869
  ultimately have "countable C" "\<forall>a\<in>C. open a"
50941
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2870
    unfolding C_def using ccover by auto
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2871
  moreover
50941
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2872
  have "\<Union>A \<inter> U \<subseteq> \<Union>C"
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2873
  proof safe
50941
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2874
    fix x a assume "x \<in> U" "x \<in> a" "a \<in> A"
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2875
    with basis[of a x] A obtain b where "b \<in> B" "x \<in> b" "b \<inter> U \<subseteq> a" by blast
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2876
    with `a \<in> A` show "x \<in> \<Union>C" unfolding C_def
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2877
      by auto
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2878
  qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2879
  then have "U \<subseteq> \<Union>C" using `U \<subseteq> \<Union>A` by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2880
  ultimately obtain T where "T\<subseteq>C" "finite T" "U \<subseteq> \<Union>T"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2881
    using * by metis
50941
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2882
  moreover then have "\<forall>t\<in>T. \<exists>a\<in>A. t \<inter> U \<subseteq> a"
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2883
    by (auto simp: C_def)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2884
  then guess f unfolding bchoice_iff Bex_def ..
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2885
  ultimately show "\<exists>T\<subseteq>A. finite T \<and> U \<subseteq> \<Union>T"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2886
    unfolding C_def by (intro exI[of _ "f`T"]) fastforce
50941
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2887
qed
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2888
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2889
lemma countably_compact_imp_compact_second_countable:
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2890
  "countably_compact U \<Longrightarrow> compact (U :: 'a :: second_countable_topology set)"
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2891
proof (rule countably_compact_imp_compact)
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2892
  fix T and x :: 'a assume "open T" "x \<in> T"
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2893
  from topological_basisE[OF is_basis this] guess b .
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2894
  then show "\<exists>b\<in>SOME B. countable B \<and> topological_basis B. x \<in> b \<and> b \<inter> U \<subseteq> T" by auto
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2895
qed (insert countable_basis topological_basis_open[OF is_basis], auto)
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  2896
50943
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  2897
lemma countably_compact_eq_compact:
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  2898
  "countably_compact U \<longleftrightarrow> compact (U :: 'a :: second_countable_topology set)"
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  2899
  using countably_compact_imp_compact_second_countable compact_imp_countably_compact by blast
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  2900
  
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  2901
subsubsection{* Sequential compactness *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2902
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2903
definition
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2904
  seq_compact :: "'a::topological_space set \<Rightarrow> bool" where
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2905
  "seq_compact S \<longleftrightarrow>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2906
   (\<forall>f. (\<forall>n. f n \<in> S) \<longrightarrow>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2907
       (\<exists>l\<in>S. \<exists>r. subseq r \<and> ((f o r) ---> l) sequentially))"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2908
50941
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2909
lemma seq_compact_imp_countably_compact:
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2910
  fixes U :: "'a :: first_countable_topology set"
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2911
  assumes "seq_compact U"
50941
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2912
  shows "countably_compact U"
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2913
proof (safe intro!: countably_compactI)
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2914
  fix A assume A: "\<forall>a\<in>A. open a" "U \<subseteq> \<Union>A" "countable A"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2915
  have subseq: "\<And>X. range X \<subseteq> U \<Longrightarrow> \<exists>r x. x \<in> U \<and> subseq r \<and> (X \<circ> r) ----> x"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2916
    using `seq_compact U` by (fastforce simp: seq_compact_def subset_eq)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2917
  show "\<exists>T\<subseteq>A. finite T \<and> U \<subseteq> \<Union>T"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2918
  proof cases
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2919
    assume "finite A" with A show ?thesis by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2920
  next
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2921
    assume "infinite A"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2922
    then have "A \<noteq> {}" by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2923
    show ?thesis
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2924
    proof (rule ccontr)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2925
      assume "\<not> (\<exists>T\<subseteq>A. finite T \<and> U \<subseteq> \<Union>T)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2926
      then have "\<forall>T. \<exists>x. T \<subseteq> A \<and> finite T \<longrightarrow> (x \<in> U - \<Union>T)" by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2927
      then obtain X' where T: "\<And>T. T \<subseteq> A \<Longrightarrow> finite T \<Longrightarrow> X' T \<in> U - \<Union>T" by metis
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2928
      def X \<equiv> "\<lambda>n. X' (from_nat_into A ` {.. n})"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2929
      have X: "\<And>n. X n \<in> U - (\<Union>i\<le>n. from_nat_into A i)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2930
        using `A \<noteq> {}` unfolding X_def SUP_def by (intro T) (auto intro: from_nat_into)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2931
      then have "range X \<subseteq> U" by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2932
      with subseq[of X] obtain r x where "x \<in> U" and r: "subseq r" "(X \<circ> r) ----> x" by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2933
      from `x\<in>U` `U \<subseteq> \<Union>A` from_nat_into_surj[OF `countable A`]
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2934
      obtain n where "x \<in> from_nat_into A n" by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2935
      with r(2) A(1) from_nat_into[OF `A \<noteq> {}`, of n]
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2936
      have "eventually (\<lambda>i. X (r i) \<in> from_nat_into A n) sequentially"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2937
        unfolding tendsto_def by (auto simp: comp_def)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2938
      then obtain N where "\<And>i. N \<le> i \<Longrightarrow> X (r i) \<in> from_nat_into A n"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2939
        by (auto simp: eventually_sequentially)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2940
      moreover from X have "\<And>i. n \<le> r i \<Longrightarrow> X (r i) \<notin> from_nat_into A n"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2941
        by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2942
      moreover from `subseq r`[THEN seq_suble, of "max n N"] have "\<exists>i. n \<le> r i \<and> N \<le> i"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2943
        by (auto intro!: exI[of _ "max n N"])
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2944
      ultimately show False
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2945
        by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2946
    qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2947
  qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2948
qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2949
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2950
lemma compact_imp_seq_compact:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2951
  fixes U :: "'a :: first_countable_topology set"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2952
  assumes "compact U" shows "seq_compact U"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2953
  unfolding seq_compact_def
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2954
proof safe
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2955
  fix X :: "nat \<Rightarrow> 'a" assume "\<forall>n. X n \<in> U"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2956
  then have "eventually (\<lambda>x. x \<in> U) (filtermap X sequentially)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2957
    by (auto simp: eventually_filtermap)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2958
  moreover have "filtermap X sequentially \<noteq> bot"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2959
    by (simp add: trivial_limit_def eventually_filtermap)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2960
  ultimately obtain x where "x \<in> U" and x: "inf (nhds x) (filtermap X sequentially) \<noteq> bot" (is "?F \<noteq> _")
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2961
    using `compact U` by (auto simp: compact_filter)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2962
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2963
  from countable_basis_at_decseq[of x] guess A . note A = this
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2964
  def s \<equiv> "\<lambda>n i. SOME j. i < j \<and> X j \<in> A (Suc n)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2965
  { fix n i
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2966
    have "\<exists>a. i < a \<and> X a \<in> A (Suc n)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2967
    proof (rule ccontr)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2968
      assume "\<not> (\<exists>a>i. X a \<in> A (Suc n))"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2969
      then have "\<And>a. Suc i \<le> a \<Longrightarrow> X a \<notin> A (Suc n)" by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2970
      then have "eventually (\<lambda>x. x \<notin> A (Suc n)) (filtermap X sequentially)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2971
        by (auto simp: eventually_filtermap eventually_sequentially)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2972
      moreover have "eventually (\<lambda>x. x \<in> A (Suc n)) (nhds x)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2973
        using A(1,2)[of "Suc n"] by (auto simp: eventually_nhds)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2974
      ultimately have "eventually (\<lambda>x. False) ?F"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2975
        by (auto simp add: eventually_inf)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2976
      with x show False
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2977
        by (simp add: eventually_False)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2978
    qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2979
    then have "i < s n i" "X (s n i) \<in> A (Suc n)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2980
      unfolding s_def by (auto intro: someI2_ex) }
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2981
  note s = this
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2982
  def r \<equiv> "nat_rec (s 0 0) s"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2983
  have "subseq r"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2984
    by (auto simp: r_def s subseq_Suc_iff)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2985
  moreover
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2986
  have "(\<lambda>n. X (r n)) ----> x"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2987
  proof (rule topological_tendstoI)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2988
    fix S assume "open S" "x \<in> S"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2989
    with A(3) have "eventually (\<lambda>i. A i \<subseteq> S) sequentially" by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2990
    moreover
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2991
    { fix i assume "Suc 0 \<le> i" then have "X (r i) \<in> A i"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2992
        by (cases i) (simp_all add: r_def s) }
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2993
    then have "eventually (\<lambda>i. X (r i) \<in> A i) sequentially" by (auto simp: eventually_sequentially)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2994
    ultimately show "eventually (\<lambda>i. X (r i) \<in> S) sequentially"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2995
      by eventually_elim auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2996
  qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2997
  ultimately show "\<exists>x \<in> U. \<exists>r. subseq r \<and> (X \<circ> r) ----> x"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2998
    using `x \<in> U` by (auto simp: convergent_def comp_def)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2999
qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  3000
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  3001
lemma seq_compactI:
44075
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3002
  assumes "\<And>f. \<forall>n. f n \<in> S \<Longrightarrow> \<exists>l\<in>S. \<exists>r. subseq r \<and> ((f o r) ---> l) sequentially"
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  3003
  shows "seq_compact S"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  3004
  unfolding seq_compact_def using assms by fast
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  3005
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  3006
lemma seq_compactE:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  3007
  assumes "seq_compact S" "\<forall>n. f n \<in> S"
44075
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3008
  obtains l r where "l \<in> S" "subseq r" "((f \<circ> r) ---> l) sequentially"
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  3009
  using assms unfolding seq_compact_def by fast
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  3010
50943
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3011
lemma countably_compact_imp_acc_point:
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3012
  assumes "countably_compact s" "countable t" "infinite t"  "t \<subseteq> s"
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3013
  shows "\<exists>x\<in>s. \<forall>U. x\<in>U \<and> open U \<longrightarrow> infinite (U \<inter> t)"
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3014
proof (rule ccontr)
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3015
  def C \<equiv> "(\<lambda>F. interior (F \<union> (- t))) ` {F. finite F \<and> F \<subseteq> t }"  
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3016
  note `countably_compact s`
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3017
  moreover have "\<forall>t\<in>C. open t" 
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3018
    by (auto simp: C_def)
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3019
  moreover
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3020
  assume "\<not> (\<exists>x\<in>s. \<forall>U. x\<in>U \<and> open U \<longrightarrow> infinite (U \<inter> t))"
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3021
  then have s: "\<And>x. x \<in> s \<Longrightarrow> \<exists>U. x\<in>U \<and> open U \<and> finite (U \<inter> t)" by metis
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3022
  have "s \<subseteq> \<Union>C"
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3023
    using `t \<subseteq> s`
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3024
    unfolding C_def Union_image_eq
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3025
    apply (safe dest!: s)
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3026
    apply (rule_tac a="U \<inter> t" in UN_I)
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3027
    apply (auto intro!: interiorI simp add: finite_subset)
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3028
    done
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3029
  moreover
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3030
  from `countable t` have "countable C"
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3031
    unfolding C_def by (auto intro: countable_Collect_finite_subset)
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3032
  ultimately guess D by (rule countably_compactE)
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3033
  then obtain E where E: "E \<subseteq> {F. finite F \<and> F \<subseteq> t }" "finite E" and
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3034
    s: "s \<subseteq> (\<Union>F\<in>E. interior (F \<union> (- t)))"
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3035
    by (metis (lifting) Union_image_eq finite_subset_image C_def)
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3036
  from s `t \<subseteq> s` have "t \<subseteq> \<Union>E"
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3037
    using interior_subset by blast
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3038
  moreover have "finite (\<Union>E)"
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3039
    using E by auto
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3040
  ultimately show False using `infinite t` by (auto simp: finite_subset)
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3041
qed
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3042
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3043
lemma countable_acc_point_imp_seq_compact:
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3044
  fixes s :: "'a::first_countable_topology set"
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3045
  assumes "\<forall>t. infinite t \<and> countable t \<and> t \<subseteq> s --> (\<exists>x\<in>s. \<forall>U. x\<in>U \<and> open U \<longrightarrow> infinite (U \<inter> t))"
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  3046
  shows "seq_compact s"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  3047
proof -
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  3048
  { fix f :: "nat \<Rightarrow> 'a" assume f: "\<forall>n. f n \<in> s"
50943
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3049
    have "\<exists>l\<in>s. \<exists>r. subseq r \<and> ((f \<circ> r) ---> l) sequentially"
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  3050
    proof (cases "finite (range f)")
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  3051
      case True
50941
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  3052
      obtain l where "infinite {n. f n = f l}"
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  3053
        using pigeonhole_infinite[OF _ True] by auto
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  3054
      then obtain r where "subseq r" and fr: "\<forall>n. f (r n) = f l"
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  3055
        using infinite_enumerate by blast
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  3056
      hence "subseq r \<and> (f \<circ> r) ----> f l"
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  3057
        by (simp add: fr tendsto_const o_def)
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  3058
      with f show "\<exists>l\<in>s. \<exists>r. subseq r \<and> (f \<circ> r) ----> l"
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  3059
        by auto
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  3060
    next
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  3061
      case False
50943
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3062
      with f assms have "\<exists>x\<in>s. \<forall>U. x\<in>U \<and> open U \<longrightarrow> infinite (U \<inter> range f)" by auto
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3063
      then obtain l where "l \<in> s" "\<forall>U. l\<in>U \<and> open U \<longrightarrow> infinite (U \<inter> range f)" ..
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3064
      from this(2) have "\<exists>r. subseq r \<and> ((f \<circ> r) ---> l) sequentially"
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3065
        using acc_point_range_imp_convergent_subsequence[of l f] by auto
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3066
      with `l \<in> s` show "\<exists>l\<in>s. \<exists>r. subseq r \<and> ((f \<circ> r) ---> l) sequentially" ..
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  3067
    qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  3068
  }
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  3069
  thus ?thesis unfolding seq_compact_def by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  3070
qed
44075
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3071
50943
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3072
lemma seq_compact_eq_countably_compact:
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3073
  fixes U :: "'a :: first_countable_topology set"
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3074
  shows "seq_compact U \<longleftrightarrow> countably_compact U"
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3075
  using
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3076
    countable_acc_point_imp_seq_compact
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3077
    countably_compact_imp_acc_point
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3078
    seq_compact_imp_countably_compact
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3079
  by metis
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3080
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3081
lemma seq_compact_eq_acc_point:
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3082
  fixes s :: "'a :: first_countable_topology set"
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3083
  shows "seq_compact s \<longleftrightarrow> (\<forall>t. infinite t \<and> countable t \<and> t \<subseteq> s --> (\<exists>x\<in>s. \<forall>U. x\<in>U \<and> open U \<longrightarrow> infinite (U \<inter> t)))"
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3084
  using
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3085
    countable_acc_point_imp_seq_compact[of s]
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3086
    countably_compact_imp_acc_point[of s]
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3087
    seq_compact_imp_countably_compact[of s]
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3088
  by metis
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3089
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3090
lemma seq_compact_eq_compact:
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3091
  fixes U :: "'a :: second_countable_topology set"
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3092
  shows "seq_compact U \<longleftrightarrow> compact U"
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3093
  using seq_compact_eq_countably_compact countably_compact_eq_compact by blast
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3094
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3095
lemma bolzano_weierstrass_imp_seq_compact:
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3096
  fixes s :: "'a::{t1_space, first_countable_topology} set"
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3097
  shows "\<forall>t. infinite t \<and> t \<subseteq> s --> (\<exists>x \<in> s. x islimpt t) \<Longrightarrow> seq_compact s"
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3098
  by (rule countable_acc_point_imp_seq_compact) (metis islimpt_eq_acc_point)
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3099
50940
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3100
subsubsection{* Total boundedness *}
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3101
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3102
lemma cauchy_def:
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3103
  "Cauchy s \<longleftrightarrow> (\<forall>e>0. \<exists>N. \<forall>m n. m \<ge> N \<and> n \<ge> N --> dist(s m)(s n) < e)"
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3104
unfolding Cauchy_def by blast
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3105
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3106
fun helper_1::"('a::metric_space set) \<Rightarrow> real \<Rightarrow> nat \<Rightarrow> 'a" where
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3107
  "helper_1 s e n = (SOME y::'a. y \<in> s \<and> (\<forall>m<n. \<not> (dist (helper_1 s e m) y < e)))"
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3108
declare helper_1.simps[simp del]
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3109
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3110
lemma seq_compact_imp_totally_bounded:
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3111
  assumes "seq_compact s"
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3112
  shows "\<forall>e>0. \<exists>k. finite k \<and> k \<subseteq> s \<and> s \<subseteq> (\<Union>((\<lambda>x. ball x e) ` k))"
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3113
proof(rule, rule, rule ccontr)
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3114
  fix e::real assume "e>0" and assm:"\<not> (\<exists>k. finite k \<and> k \<subseteq> s \<and> s \<subseteq> \<Union>(\<lambda>x. ball x e) ` k)"
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3115
  def x \<equiv> "helper_1 s e"
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3116
  { fix n
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3117
    have "x n \<in> s \<and> (\<forall>m<n. \<not> dist (x m) (x n) < e)"
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3118
    proof(induct_tac rule:nat_less_induct)
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3119
      fix n  def Q \<equiv> "(\<lambda>y. y \<in> s \<and> (\<forall>m<n. \<not> dist (x m) y < e))"
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3120
      assume as:"\<forall>m<n. x m \<in> s \<and> (\<forall>ma<m. \<not> dist (x ma) (x m) < e)"
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3121
      have "\<not> s \<subseteq> (\<Union>x\<in>x ` {0..<n}. ball x e)" using assm apply simp apply(erule_tac x="x ` {0 ..< n}" in allE) using as by auto
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3122
      then obtain z where z:"z\<in>s" "z \<notin> (\<Union>x\<in>x ` {0..<n}. ball x e)" unfolding subset_eq by auto
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3123
      have "Q (x n)" unfolding x_def and helper_1.simps[of s e n]
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3124
        apply(rule someI2[where a=z]) unfolding x_def[symmetric] and Q_def using z by auto
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3125
      thus "x n \<in> s \<and> (\<forall>m<n. \<not> dist (x m) (x n) < e)" unfolding Q_def by auto
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3126
    qed }
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3127
  hence "\<forall>n::nat. x n \<in> s" and x:"\<forall>n. \<forall>m < n. \<not> (dist (x m) (x n) < e)" by blast+
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3128
  then obtain l r where "l\<in>s" and r:"subseq r" and "((x \<circ> r) ---> l) sequentially" using assms(1)[unfolded seq_compact_def, THEN spec[where x=x]] by auto
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3129
  from this(3) have "Cauchy (x \<circ> r)" using LIMSEQ_imp_Cauchy by auto
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3130
  then obtain N::nat where N:"\<forall>m n. N \<le> m \<and> N \<le> n \<longrightarrow> dist ((x \<circ> r) m) ((x \<circ> r) n) < e" unfolding cauchy_def using `e>0` by auto
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3131
  show False
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3132
    using N[THEN spec[where x=N], THEN spec[where x="N+1"]]
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3133
    using r[unfolded subseq_def, THEN spec[where x=N], THEN spec[where x="N+1"]]
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3134
    using x[THEN spec[where x="r (N+1)"], THEN spec[where x="r (N)"]] by auto
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3135
qed
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3136
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3137
subsubsection{* Heine-Borel theorem *}
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3138
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3139
lemma seq_compact_imp_heine_borel:
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3140
  fixes s :: "'a :: metric_space set"
50943
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3141
  assumes "seq_compact s" shows "compact s"
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3142
proof -
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3143
  from seq_compact_imp_totally_bounded[OF `seq_compact s`]
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3144
  guess f unfolding choice_iff' .. note f = this
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3145
  def K \<equiv> "(\<lambda>(x, r). ball x r) ` ((\<Union>e \<in> \<rat> \<inter> {0 <..}. f e) \<times> \<rat>)"
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3146
  have "countably_compact s"
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3147
    using `seq_compact s` by (rule seq_compact_imp_countably_compact)
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3148
  then show "compact s"
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3149
  proof (rule countably_compact_imp_compact)
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3150
    show "countable K"
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3151
      unfolding K_def using f
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3152
      by (auto intro: countable_finite countable_subset countable_rat
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3153
               intro!: countable_image countable_SIGMA countable_UN)
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3154
    show "\<forall>b\<in>K. open b" by (auto simp: K_def)
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3155
  next
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3156
    fix T x assume T: "open T" "x \<in> T" and x: "x \<in> s"
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3157
    from openE[OF T] obtain e where "0 < e" "ball x e \<subseteq> T" by auto
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3158
    then have "0 < e / 2" "ball x (e / 2) \<subseteq> T" by auto
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3159
    from Rats_dense_in_real[OF `0 < e / 2`] obtain r where "r \<in> \<rat>" "0 < r" "r < e / 2" by auto
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3160
    from f[rule_format, of r] `0 < r` `x \<in> s` obtain k where "k \<in> f r" "x \<in> ball k r"
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3161
      unfolding Union_image_eq by auto
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3162
    from `r \<in> \<rat>` `0 < r` `k \<in> f r` have "ball k r \<in> K" by (auto simp: K_def)
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3163
    then show "\<exists>b\<in>K. x \<in> b \<and> b \<inter> s \<subseteq> T"
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3164
    proof (rule bexI[rotated], safe)
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3165
      fix y assume "y \<in> ball k r"
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3166
      with `r < e / 2` `x \<in> ball k r` have "dist x y < e"
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3167
        by (intro dist_double[where x = k and d=e]) (auto simp: dist_commute)
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3168
      with `ball x e \<subseteq> T` show "y \<in> T" by auto
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3169
    qed (rule `x \<in> ball k r`)
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3170
  qed
50940
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3171
qed
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3172
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3173
lemma compact_eq_seq_compact_metric:
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3174
  "compact (s :: 'a::metric_space set) \<longleftrightarrow> seq_compact s"
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3175
  using compact_imp_seq_compact seq_compact_imp_heine_borel by blast
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3176
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3177
lemma compact_def:
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3178
  "compact (S :: 'a::metric_space set) \<longleftrightarrow>
50943
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  3179
   (\<forall>f. (\<forall>n. f n \<in> S) \<longrightarrow> (\<exists>l\<in>S. \<exists>r. subseq r \<and> (f o r) ----> l))"
50940
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3180
  unfolding compact_eq_seq_compact_metric seq_compact_def by auto
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3181
50944
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  3182
subsubsection {* Complete the chain of compactness variants *}
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  3183
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  3184
lemma compact_eq_bolzano_weierstrass:
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  3185
  fixes s :: "'a::metric_space set"
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  3186
  shows "compact s \<longleftrightarrow> (\<forall>t. infinite t \<and> t \<subseteq> s --> (\<exists>x \<in> s. x islimpt t))" (is "?lhs = ?rhs")
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  3187
proof
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  3188
  assume ?lhs thus ?rhs using heine_borel_imp_bolzano_weierstrass[of s] by auto
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  3189
next
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  3190
  assume ?rhs thus ?lhs
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  3191
    unfolding compact_eq_seq_compact_metric by (rule bolzano_weierstrass_imp_seq_compact)
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  3192
qed
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  3193
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  3194
lemma bolzano_weierstrass_imp_bounded:
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  3195
  "\<forall>t. infinite t \<and> t \<subseteq> s --> (\<exists>x \<in> s. x islimpt t) \<Longrightarrow> bounded s"
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  3196
  using compact_imp_bounded unfolding compact_eq_bolzano_weierstrass .
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  3197
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3198
text {*
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3199
  A metric space (or topological vector space) is said to have the
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3200
  Heine-Borel property if every closed and bounded subset is compact.
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3201
*}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3202
44207
ea99698c2070 locale-ize some definitions, so perfect_space and heine_borel can inherit from the proper superclasses
huffman
parents: 44170
diff changeset
  3203
class heine_borel = metric_space +
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3204
  assumes bounded_imp_convergent_subsequence:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3205
    "bounded s \<Longrightarrow> \<forall>n. f n \<in> s
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3206
      \<Longrightarrow> \<exists>l r. subseq r \<and> ((f \<circ> r) ---> l) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3207
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  3208
lemma bounded_closed_imp_seq_compact:
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3209
  fixes s::"'a::heine_borel set"
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  3210
  assumes "bounded s" and "closed s" shows "seq_compact s"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  3211
proof (unfold seq_compact_def, clarify)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3212
  fix f :: "nat \<Rightarrow> 'a" assume f: "\<forall>n. f n \<in> s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3213
  obtain l r where r: "subseq r" and l: "((f \<circ> r) ---> l) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3214
    using bounded_imp_convergent_subsequence [OF `bounded s` `\<forall>n. f n \<in> s`] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3215
  from f have fr: "\<forall>n. (f \<circ> r) n \<in> s" by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3216
  have "l \<in> s" using `closed s` fr l
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3217
    unfolding closed_sequential_limits by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3218
  show "\<exists>l\<in>s. \<exists>r. subseq r \<and> ((f \<circ> r) ---> l) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3219
    using `l \<in> s` r l by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3220
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3221
50944
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  3222
lemma compact_eq_bounded_closed:
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  3223
  fixes s :: "'a::heine_borel set"
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  3224
  shows "compact s \<longleftrightarrow> bounded s \<and> closed s"  (is "?lhs = ?rhs")
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  3225
proof
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  3226
  assume ?lhs thus ?rhs
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  3227
    using compact_imp_closed compact_imp_bounded by blast
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  3228
next
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  3229
  assume ?rhs thus ?lhs
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  3230
    using bounded_closed_imp_seq_compact[of s] unfolding compact_eq_seq_compact_metric by auto
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  3231
qed
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  3232
50973
4a2c82644889 generalized diameter from real_normed_vector to metric_space
hoelzl
parents: 50972
diff changeset
  3233
(* TODO: is this lemma necessary? *)
50972
d2c6a0a7fcdf tuned proof
hoelzl
parents: 50971
diff changeset
  3234
lemma bounded_increasing_convergent:
d2c6a0a7fcdf tuned proof
hoelzl
parents: 50971
diff changeset
  3235
  fixes s :: "nat \<Rightarrow> real"
50973
4a2c82644889 generalized diameter from real_normed_vector to metric_space
hoelzl
parents: 50972
diff changeset
  3236
  shows "bounded {s n| n. True} \<Longrightarrow> \<forall>n. s n \<le> s (Suc n) \<Longrightarrow> \<exists>l. s ----> l"
50972
d2c6a0a7fcdf tuned proof
hoelzl
parents: 50971
diff changeset
  3237
  using Bseq_mono_convergent[of s] incseq_Suc_iff[of s]
d2c6a0a7fcdf tuned proof
hoelzl
parents: 50971
diff changeset
  3238
  by (auto simp: image_def Bseq_eq_bounded convergent_def incseq_def)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3239
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3240
instance real :: heine_borel
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3241
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3242
  fix s :: "real set" and f :: "nat \<Rightarrow> real"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3243
  assume s: "bounded s" and f: "\<forall>n. f n \<in> s"
50972
d2c6a0a7fcdf tuned proof
hoelzl
parents: 50971
diff changeset
  3244
  obtain r where r: "subseq r" "monoseq (f \<circ> r)"
d2c6a0a7fcdf tuned proof
hoelzl
parents: 50971
diff changeset
  3245
    unfolding comp_def by (metis seq_monosub)
d2c6a0a7fcdf tuned proof
hoelzl
parents: 50971
diff changeset
  3246
  moreover
d2c6a0a7fcdf tuned proof
hoelzl
parents: 50971
diff changeset
  3247
  then have "Bseq (f \<circ> r)"
d2c6a0a7fcdf tuned proof
hoelzl
parents: 50971
diff changeset
  3248
    unfolding Bseq_eq_bounded using s f by (auto intro: bounded_subset)
d2c6a0a7fcdf tuned proof
hoelzl
parents: 50971
diff changeset
  3249
  ultimately show "\<exists>l r. subseq r \<and> (f \<circ> r) ----> l"
d2c6a0a7fcdf tuned proof
hoelzl
parents: 50971
diff changeset
  3250
    using Bseq_monoseq_convergent[of "f \<circ> r"] by (auto simp: convergent_def)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3251
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3252
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3253
lemma compact_lemma:
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  3254
  fixes f :: "nat \<Rightarrow> 'a::euclidean_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3255
  assumes "bounded s" and "\<forall>n. f n \<in> s"
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  3256
  shows "\<forall>d\<subseteq>Basis. \<exists>l::'a. \<exists> r. subseq r \<and>
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  3257
        (\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r n) \<bullet> i) (l \<bullet> i) < e) sequentially)"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  3258
proof safe
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  3259
  fix d :: "'a set" assume d: "d \<subseteq> Basis" 
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  3260
  with finite_Basis have "finite d" by (blast intro: finite_subset)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  3261
  from this d show "\<exists>l::'a. \<exists>r. subseq r \<and>
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  3262
      (\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r n) \<bullet> i) (l \<bullet> i) < e) sequentially)"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3263
  proof(induct d) case empty thus ?case unfolding subseq_def by auto
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  3264
  next case (insert k d) have k[intro]:"k\<in>Basis" using insert by auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  3265
    have s': "bounded ((\<lambda>x. x \<bullet> k) ` s)" using `bounded s`
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  3266
      by (auto intro!: bounded_linear_image bounded_linear_inner_left)
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  3267
    obtain l1::"'a" and r1 where r1:"subseq r1" and
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  3268
      lr1:"\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r1 n) \<bullet> i) (l1 \<bullet> i) < e) sequentially"
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  3269
      using insert(3) using insert(4) by auto
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  3270
    have f': "\<forall>n. f (r1 n) \<bullet> k \<in> (\<lambda>x. x \<bullet> k) ` s" using `\<forall>n. f n \<in> s` by simp
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  3271
    obtain l2 r2 where r2:"subseq r2" and lr2:"((\<lambda>i. f (r1 (r2 i)) \<bullet> k) ---> l2) sequentially"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3272
      using bounded_imp_convergent_subsequence[OF s' f'] unfolding o_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3273
    def r \<equiv> "r1 \<circ> r2" have r:"subseq r"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3274
      using r1 and r2 unfolding r_def o_def subseq_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3275
    moreover
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  3276
    def l \<equiv> "(\<Sum>i\<in>Basis. (if i = k then l2 else l1\<bullet>i) *\<^sub>R i)::'a"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3277
    { fix e::real assume "e>0"
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  3278
      from lr1 `e>0` have N1:"eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r1 n) \<bullet> i) (l1 \<bullet> i) < e) sequentially" by blast
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  3279
      from lr2 `e>0` have N2:"eventually (\<lambda>n. dist (f (r1 (r2 n)) \<bullet> k) l2 < e) sequentially" by (rule tendstoD)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  3280
      from r2 N1 have N1': "eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r1 (r2 n)) \<bullet> i) (l1 \<bullet> i) < e) sequentially"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3281
        by (rule eventually_subseq)
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  3282
      have "eventually (\<lambda>n. \<forall>i\<in>(insert k d). dist (f (r n) \<bullet> i) (l \<bullet> i) < e) sequentially"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  3283
        using N1' N2 
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  3284
        by eventually_elim (insert insert.prems, auto simp: l_def r_def o_def)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3285
    }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3286
    ultimately show ?case by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3287
  qed
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  3288
qed
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  3289
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  3290
instance euclidean_space \<subseteq> heine_borel
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3291
proof
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  3292
  fix s :: "'a set" and f :: "nat \<Rightarrow> 'a"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3293
  assume s: "bounded s" and f: "\<forall>n. f n \<in> s"
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  3294
  then obtain l::'a and r where r: "subseq r"
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  3295
    and l: "\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>Basis. dist (f (r n) \<bullet> i) (l \<bullet> i) < e) sequentially"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3296
    using compact_lemma [OF s f] by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3297
  { fix e::real assume "e>0"
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  3298
    hence "0 < e / real_of_nat DIM('a)" by (auto intro!: divide_pos_pos DIM_positive)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  3299
    with l have "eventually (\<lambda>n. \<forall>i\<in>Basis. dist (f (r n) \<bullet> i) (l \<bullet> i) < e / (real_of_nat DIM('a))) sequentially"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3300
      by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3301
    moreover
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  3302
    { fix n assume n: "\<forall>i\<in>Basis. dist (f (r n) \<bullet> i) (l \<bullet> i) < e / (real_of_nat DIM('a))"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  3303
      have "dist (f (r n)) l \<le> (\<Sum>i\<in>Basis. dist (f (r n) \<bullet> i) (l \<bullet> i))"
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  3304
        apply(subst euclidean_dist_l2) using zero_le_dist by (rule setL2_le_setsum)
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  3305
      also have "\<dots> < (\<Sum>i\<in>(Basis::'a set). e / (real_of_nat DIM('a)))"
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  3306
        apply(rule setsum_strict_mono) using n by auto
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  3307
      finally have "dist (f (r n)) l < e" 
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  3308
        by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3309
    }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3310
    ultimately have "eventually (\<lambda>n. dist (f (r n)) l < e) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3311
      by (rule eventually_elim1)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3312
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3313
  hence *:"((f \<circ> r) ---> l) sequentially" unfolding o_def tendsto_iff by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3314
  with r show "\<exists>l r. subseq r \<and> ((f \<circ> r) ---> l) sequentially" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3315
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3316
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3317
lemma bounded_fst: "bounded s \<Longrightarrow> bounded (fst ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3318
unfolding bounded_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3319
apply clarify
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3320
apply (rule_tac x="a" in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3321
apply (rule_tac x="e" in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3322
apply clarsimp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3323
apply (drule (1) bspec)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3324
apply (simp add: dist_Pair_Pair)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3325
apply (erule order_trans [OF real_sqrt_sum_squares_ge1])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3326
done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3327
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3328
lemma bounded_snd: "bounded s \<Longrightarrow> bounded (snd ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3329
unfolding bounded_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3330
apply clarify
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3331
apply (rule_tac x="b" in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3332
apply (rule_tac x="e" in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3333
apply clarsimp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3334
apply (drule (1) bspec)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3335
apply (simp add: dist_Pair_Pair)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3336
apply (erule order_trans [OF real_sqrt_sum_squares_ge2])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3337
done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3338
37678
0040bafffdef "prod" and "sum" replace "*" and "+" respectively
haftmann
parents: 37649
diff changeset
  3339
instance prod :: (heine_borel, heine_borel) heine_borel
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3340
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3341
  fix s :: "('a * 'b) set" and f :: "nat \<Rightarrow> 'a * 'b"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3342
  assume s: "bounded s" and f: "\<forall>n. f n \<in> s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3343
  from s have s1: "bounded (fst ` s)" by (rule bounded_fst)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3344
  from f have f1: "\<forall>n. fst (f n) \<in> fst ` s" by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3345
  obtain l1 r1 where r1: "subseq r1"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3346
    and l1: "((\<lambda>n. fst (f (r1 n))) ---> l1) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3347
    using bounded_imp_convergent_subsequence [OF s1 f1]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3348
    unfolding o_def by fast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3349
  from s have s2: "bounded (snd ` s)" by (rule bounded_snd)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3350
  from f have f2: "\<forall>n. snd (f (r1 n)) \<in> snd ` s" by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3351
  obtain l2 r2 where r2: "subseq r2"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3352
    and l2: "((\<lambda>n. snd (f (r1 (r2 n)))) ---> l2) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3353
    using bounded_imp_convergent_subsequence [OF s2 f2]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3354
    unfolding o_def by fast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3355
  have l1': "((\<lambda>n. fst (f (r1 (r2 n)))) ---> l1) sequentially"
50972
d2c6a0a7fcdf tuned proof
hoelzl
parents: 50971
diff changeset
  3356
    using LIMSEQ_subseq_LIMSEQ [OF l1 r2] unfolding o_def .
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3357
  have l: "((f \<circ> (r1 \<circ> r2)) ---> (l1, l2)) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3358
    using tendsto_Pair [OF l1' l2] unfolding o_def by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3359
  have r: "subseq (r1 \<circ> r2)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3360
    using r1 r2 unfolding subseq_def by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3361
  show "\<exists>l r. subseq r \<and> ((f \<circ> r) ---> l) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3362
    using l r by fast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3363
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3364
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  3365
subsubsection{* Completeness *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3366
50971
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3367
definition complete :: "'a::metric_space set \<Rightarrow> bool" where
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3368
  "complete s \<longleftrightarrow> (\<forall>f. (\<forall>n. f n \<in> s) \<and> Cauchy f \<longrightarrow> (\<exists>l\<in>s. f ----> l))"
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3369
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3370
lemma compact_imp_complete: assumes "compact s" shows "complete s"
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3371
proof-
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3372
  { fix f assume as: "(\<forall>n::nat. f n \<in> s)" "Cauchy f"
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3373
    from as(1) obtain l r where lr: "l\<in>s" "subseq r" "(f \<circ> r) ----> l"
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3374
      using assms unfolding compact_def by blast
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3375
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3376
    note lr' = seq_suble [OF lr(2)]
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3377
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3378
    { fix e::real assume "e>0"
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3379
      from as(2) obtain N where N:"\<forall>m n. N \<le> m \<and> N \<le> n \<longrightarrow> dist (f m) (f n) < e/2" unfolding cauchy_def using `e>0` apply (erule_tac x="e/2" in allE) by auto
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3380
      from lr(3)[unfolded LIMSEQ_def, THEN spec[where x="e/2"]] obtain M where M:"\<forall>n\<ge>M. dist ((f \<circ> r) n) l < e/2" using `e>0` by auto
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3381
      { fix n::nat assume n:"n \<ge> max N M"
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3382
        have "dist ((f \<circ> r) n) l < e/2" using n M by auto
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3383
        moreover have "r n \<ge> N" using lr'[of n] n by auto
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3384
        hence "dist (f n) ((f \<circ> r) n) < e / 2" using N using n by auto
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3385
        ultimately have "dist (f n) l < e" using dist_triangle_half_r[of "f (r n)" "f n" e l] by (auto simp add: dist_commute)  }
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3386
      hence "\<exists>N. \<forall>n\<ge>N. dist (f n) l < e" by blast  }
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3387
    hence "\<exists>l\<in>s. (f ---> l) sequentially" using `l\<in>s` unfolding LIMSEQ_def by auto  }
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3388
  thus ?thesis unfolding complete_def by auto
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3389
qed
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3390
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3391
lemma nat_approx_posE:
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3392
  fixes e::real
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3393
  assumes "0 < e"
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3394
  obtains n::nat where "1 / (Suc n) < e"
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3395
proof atomize_elim
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3396
  have " 1 / real (Suc (nat (ceiling (1/e)))) < 1 / (ceiling (1/e))"
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3397
    by (rule divide_strict_left_mono) (auto intro!: mult_pos_pos simp: `0 < e`)
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3398
  also have "1 / (ceiling (1/e)) \<le> 1 / (1/e)"
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3399
    by (rule divide_left_mono) (auto intro!: divide_pos_pos simp: `0 < e`)
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3400
  also have "\<dots> = e" by simp
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3401
  finally show  "\<exists>n. 1 / real (Suc n) < e" ..
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3402
qed
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3403
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3404
lemma compact_eq_totally_bounded:
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3405
  "compact s \<longleftrightarrow> complete s \<and> (\<forall>e>0. \<exists>k. finite k \<and> s \<subseteq> (\<Union>((\<lambda>x. ball x e) ` k)))"
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3406
    (is "_ \<longleftrightarrow> ?rhs")
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3407
proof
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3408
  assume assms: "?rhs"
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3409
  then obtain k where k: "\<And>e. 0 < e \<Longrightarrow> finite (k e)" "\<And>e. 0 < e \<Longrightarrow> s \<subseteq> (\<Union>x\<in>k e. ball x e)"
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3410
    by (auto simp: choice_iff')
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3411
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3412
  show "compact s"
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3413
  proof cases
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3414
    assume "s = {}" thus "compact s" by (simp add: compact_def)
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3415
  next
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3416
    assume "s \<noteq> {}"
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3417
    show ?thesis
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3418
      unfolding compact_def
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3419
    proof safe
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3420
      fix f :: "nat \<Rightarrow> 'a" assume f: "\<forall>n. f n \<in> s"
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3421
      
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3422
      def e \<equiv> "\<lambda>n. 1 / (2 * Suc n)"
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3423
      then have [simp]: "\<And>n. 0 < e n" by auto
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3424
      def B \<equiv> "\<lambda>n U. SOME b. infinite {n. f n \<in> b} \<and> (\<exists>x. b \<subseteq> ball x (e n) \<inter> U)"
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3425
      { fix n U assume "infinite {n. f n \<in> U}"
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3426
        then have "\<exists>b\<in>k (e n). infinite {i\<in>{n. f n \<in> U}. f i \<in> ball b (e n)}"
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3427
          using k f by (intro pigeonhole_infinite_rel) (auto simp: subset_eq)
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3428
        then guess a ..
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3429
        then have "\<exists>b. infinite {i. f i \<in> b} \<and> (\<exists>x. b \<subseteq> ball x (e n) \<inter> U)"
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3430
          by (intro exI[of _ "ball a (e n) \<inter> U"] exI[of _ a]) (auto simp: ac_simps)
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3431
        from someI_ex[OF this]
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3432
        have "infinite {i. f i \<in> B n U}" "\<exists>x. B n U \<subseteq> ball x (e n) \<inter> U"
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3433
          unfolding B_def by auto }
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3434
      note B = this
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3435
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3436
      def F \<equiv> "nat_rec (B 0 UNIV) B"
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3437
      { fix n have "infinite {i. f i \<in> F n}"
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3438
          by (induct n) (auto simp: F_def B) }
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3439
      then have F: "\<And>n. \<exists>x. F (Suc n) \<subseteq> ball x (e n) \<inter> F n"
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3440
        using B by (simp add: F_def)
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3441
      then have F_dec: "\<And>m n. m \<le> n \<Longrightarrow> F n \<subseteq> F m"
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3442
        using decseq_SucI[of F] by (auto simp: decseq_def)
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3443
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3444
      obtain sel where sel: "\<And>k i. i < sel k i" "\<And>k i. f (sel k i) \<in> F k"
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3445
      proof (atomize_elim, unfold all_conj_distrib[symmetric], intro choice allI)
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3446
        fix k i
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3447
        have "infinite ({n. f n \<in> F k} - {.. i})"
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3448
          using `infinite {n. f n \<in> F k}` by auto
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3449
        from infinite_imp_nonempty[OF this]
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3450
        show "\<exists>x>i. f x \<in> F k"
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3451
          by (simp add: set_eq_iff not_le conj_commute)
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3452
      qed
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3453
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3454
      def t \<equiv> "nat_rec (sel 0 0) (\<lambda>n i. sel (Suc n) i)"
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3455
      have "subseq t"
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3456
        unfolding subseq_Suc_iff by (simp add: t_def sel)
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3457
      moreover have "\<forall>i. (f \<circ> t) i \<in> s"
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3458
        using f by auto
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3459
      moreover
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3460
      { fix n have "(f \<circ> t) n \<in> F n"
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3461
          by (cases n) (simp_all add: t_def sel) }
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3462
      note t = this
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3463
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3464
      have "Cauchy (f \<circ> t)"
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3465
      proof (safe intro!: metric_CauchyI exI elim!: nat_approx_posE)
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3466
        fix r :: real and N n m assume "1 / Suc N < r" "Suc N \<le> n" "Suc N \<le> m"
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3467
        then have "(f \<circ> t) n \<in> F (Suc N)" "(f \<circ> t) m \<in> F (Suc N)" "2 * e N < r"
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3468
          using F_dec t by (auto simp: e_def field_simps real_of_nat_Suc)
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3469
        with F[of N] obtain x where "dist x ((f \<circ> t) n) < e N" "dist x ((f \<circ> t) m) < e N"
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3470
          by (auto simp: subset_eq)
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3471
        with dist_triangle[of "(f \<circ> t) m" "(f \<circ> t) n" x] `2 * e N < r`
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3472
        show "dist ((f \<circ> t) m) ((f \<circ> t) n) < r"
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3473
          by (simp add: dist_commute)
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3474
      qed
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3475
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3476
      ultimately show "\<exists>l\<in>s. \<exists>r. subseq r \<and> (f \<circ> r) ----> l"
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3477
        using assms unfolding complete_def by blast
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3478
    qed
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3479
  qed
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3480
qed (metis compact_imp_complete compact_imp_seq_compact seq_compact_imp_totally_bounded)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3481
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3482
lemma cauchy: "Cauchy s \<longleftrightarrow> (\<forall>e>0.\<exists> N::nat. \<forall>n\<ge>N. dist(s n)(s N) < e)" (is "?lhs = ?rhs")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3483
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3484
  { assume ?rhs
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3485
    { fix e::real
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3486
      assume "e>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3487
      with `?rhs` obtain N where N:"\<forall>n\<ge>N. dist (s n) (s N) < e/2"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3488
        by (erule_tac x="e/2" in allE) auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3489
      { fix n m
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3490
        assume nm:"N \<le> m \<and> N \<le> n"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3491
        hence "dist (s m) (s n) < e" using N
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3492
          using dist_triangle_half_l[of "s m" "s N" "e" "s n"]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3493
          by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3494
      }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3495
      hence "\<exists>N. \<forall>m n. N \<le> m \<and> N \<le> n \<longrightarrow> dist (s m) (s n) < e"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3496
        by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3497
    }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3498
    hence ?lhs
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3499
      unfolding cauchy_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3500
      by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3501
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3502
  thus ?thesis
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3503
    unfolding cauchy_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3504
    using dist_triangle_half_l
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3505
    by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3506
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3507
34104
22758f95e624 re-state lemmas using 'range'
huffman
parents: 33758
diff changeset
  3508
lemma cauchy_imp_bounded: assumes "Cauchy s" shows "bounded (range s)"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3509
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3510
  from assms obtain N::nat where "\<forall>m n. N \<le> m \<and> N \<le> n \<longrightarrow> dist (s m) (s n) < 1" unfolding cauchy_def apply(erule_tac x= 1 in allE) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3511
  hence N:"\<forall>n. N \<le> n \<longrightarrow> dist (s N) (s n) < 1" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3512
  moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3513
  have "bounded (s ` {0..N})" using finite_imp_bounded[of "s ` {1..N}"] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3514
  then obtain a where a:"\<forall>x\<in>s ` {0..N}. dist (s N) x \<le> a"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3515
    unfolding bounded_any_center [where a="s N"] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3516
  ultimately show "?thesis"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3517
    unfolding bounded_any_center [where a="s N"]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3518
    apply(rule_tac x="max a 1" in exI) apply auto
34104
22758f95e624 re-state lemmas using 'range'
huffman
parents: 33758
diff changeset
  3519
    apply(erule_tac x=y in allE) apply(erule_tac x=y in ballE) by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3520
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3521
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3522
instance heine_borel < complete_space
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3523
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3524
  fix f :: "nat \<Rightarrow> 'a" assume "Cauchy f"
34104
22758f95e624 re-state lemmas using 'range'
huffman
parents: 33758
diff changeset
  3525
  hence "bounded (range f)"
22758f95e624 re-state lemmas using 'range'
huffman
parents: 33758
diff changeset
  3526
    by (rule cauchy_imp_bounded)
50971
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3527
  hence "compact (closure (range f))"
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3528
    unfolding compact_eq_bounded_closed by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3529
  hence "complete (closure (range f))"
50971
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50970
diff changeset
  3530
    by (rule compact_imp_complete)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3531
  moreover have "\<forall>n. f n \<in> closure (range f)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3532
    using closure_subset [of "range f"] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3533
  ultimately have "\<exists>l\<in>closure (range f). (f ---> l) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3534
    using `Cauchy f` unfolding complete_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3535
  then show "convergent f"
36660
1cc4ab4b7ff7 make (X ----> L) an abbreviation for (X ---> L) sequentially
huffman
parents: 36659
diff changeset
  3536
    unfolding convergent_def by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3537
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3538
44632
076a45f65e12 simplify/generalize some proofs
huffman
parents: 44628
diff changeset
  3539
instance euclidean_space \<subseteq> banach ..
076a45f65e12 simplify/generalize some proofs
huffman
parents: 44628
diff changeset
  3540
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3541
lemma complete_univ: "complete (UNIV :: 'a::complete_space set)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3542
proof(simp add: complete_def, rule, rule)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3543
  fix f :: "nat \<Rightarrow> 'a" assume "Cauchy f"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3544
  hence "convergent f" by (rule Cauchy_convergent)
36660
1cc4ab4b7ff7 make (X ----> L) an abbreviation for (X ---> L) sequentially
huffman
parents: 36659
diff changeset
  3545
  thus "\<exists>l. f ----> l" unfolding convergent_def .  
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3546
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3547
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3548
lemma complete_imp_closed: assumes "complete s" shows "closed s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3549
proof -
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3550
  { fix x assume "x islimpt s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3551
    then obtain f where f: "\<forall>n. f n \<in> s - {x}" "(f ---> x) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3552
      unfolding islimpt_sequential by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3553
    then obtain l where l: "l\<in>s" "(f ---> l) sequentially"
50939
ae7cd20ed118 replace convergent_imp_cauchy by LIMSEQ_imp_Cauchy
hoelzl
parents: 50938
diff changeset
  3554
      using `complete s`[unfolded complete_def] using LIMSEQ_imp_Cauchy[of f x] by auto
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41969
diff changeset
  3555
    hence "x \<in> s"  using tendsto_unique[of sequentially f l x] trivial_limit_sequentially f(2) by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3556
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3557
  thus "closed s" unfolding closed_limpt by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3558
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3559
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3560
lemma complete_eq_closed:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3561
  fixes s :: "'a::complete_space set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3562
  shows "complete s \<longleftrightarrow> closed s" (is "?lhs = ?rhs")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3563
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3564
  assume ?lhs thus ?rhs by (rule complete_imp_closed)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3565
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3566
  assume ?rhs
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3567
  { fix f assume as:"\<forall>n::nat. f n \<in> s" "Cauchy f"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3568
    then obtain l where "(f ---> l) sequentially" using complete_univ[unfolded complete_def, THEN spec[where x=f]] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3569
    hence "\<exists>l\<in>s. (f ---> l) sequentially" using `?rhs`[unfolded closed_sequential_limits, THEN spec[where x=f], THEN spec[where x=l]] using as(1) by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3570
  thus ?lhs unfolding complete_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3571
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3572
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3573
lemma convergent_eq_cauchy:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3574
  fixes s :: "nat \<Rightarrow> 'a::complete_space"
44632
076a45f65e12 simplify/generalize some proofs
huffman
parents: 44628
diff changeset
  3575
  shows "(\<exists>l. (s ---> l) sequentially) \<longleftrightarrow> Cauchy s"
076a45f65e12 simplify/generalize some proofs
huffman
parents: 44628
diff changeset
  3576
  unfolding Cauchy_convergent_iff convergent_def ..
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3577
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3578
lemma convergent_imp_bounded:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3579
  fixes s :: "nat \<Rightarrow> 'a::metric_space"
44632
076a45f65e12 simplify/generalize some proofs
huffman
parents: 44628
diff changeset
  3580
  shows "(s ---> l) sequentially \<Longrightarrow> bounded (range s)"
50939
ae7cd20ed118 replace convergent_imp_cauchy by LIMSEQ_imp_Cauchy
hoelzl
parents: 50938
diff changeset
  3581
  by (intro cauchy_imp_bounded LIMSEQ_imp_Cauchy)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3582
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3583
lemma compact_cball[simp]:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3584
  fixes x :: "'a::heine_borel"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3585
  shows "compact(cball x e)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3586
  using compact_eq_bounded_closed bounded_cball closed_cball
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3587
  by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3588
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3589
lemma compact_frontier_bounded[intro]:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3590
  fixes s :: "'a::heine_borel set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3591
  shows "bounded s ==> compact(frontier s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3592
  unfolding frontier_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3593
  using compact_eq_bounded_closed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3594
  by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3595
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3596
lemma compact_frontier[intro]:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3597
  fixes s :: "'a::heine_borel set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3598
  shows "compact s ==> compact (frontier s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3599
  using compact_eq_bounded_closed compact_frontier_bounded
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3600
  by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3601
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3602
lemma frontier_subset_compact:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3603
  fixes s :: "'a::heine_borel set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3604
  shows "compact s ==> frontier s \<subseteq> s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3605
  using frontier_subset_closed compact_eq_bounded_closed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3606
  by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3607
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  3608
subsection {* Bounded closed nest property (proof does not use Heine-Borel) *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3609
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3610
lemma bounded_closed_nest:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3611
  assumes "\<forall>n. closed(s n)" "\<forall>n. (s n \<noteq> {})"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3612
  "(\<forall>m n. m \<le> n --> s n \<subseteq> s m)"  "bounded(s 0)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3613
  shows "\<exists>a::'a::heine_borel. \<forall>n::nat. a \<in> s(n)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3614
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3615
  from assms(2) obtain x where x:"\<forall>n::nat. x n \<in> s n" using choice[of "\<lambda>n x. x\<in> s n"] by auto
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  3616
  from assms(4,1) have *:"seq_compact (s 0)" using bounded_closed_imp_seq_compact[of "s 0"] by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3617
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3618
  then obtain l r where lr:"l\<in>s 0" "subseq r" "((x \<circ> r) ---> l) sequentially"
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  3619
    unfolding seq_compact_def apply(erule_tac x=x in allE)  using x using assms(3) by blast
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3620
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3621
  { fix n::nat
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3622
    { fix e::real assume "e>0"
44907
93943da0a010 remove redundant lemma Lim_sequentially in favor of lemma LIMSEQ_def
huffman
parents: 44905
diff changeset
  3623
      with lr(3) obtain N where N:"\<forall>m\<ge>N. dist ((x \<circ> r) m) l < e" unfolding LIMSEQ_def by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3624
      hence "dist ((x \<circ> r) (max N n)) l < e" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3625
      moreover
50937
d249ef928ae1 removed subseq_bigger (replaced by seq_suble)
hoelzl
parents: 50936
diff changeset
  3626
      have "r (max N n) \<ge> n" using lr(2) using seq_suble[of r "max N n"] by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3627
      hence "(x \<circ> r) (max N n) \<in> s n"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3628
        using x apply(erule_tac x=n in allE)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3629
        using x apply(erule_tac x="r (max N n)" in allE)
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  3630
        using assms(3) apply(erule_tac x=n in allE) apply(erule_tac x="r (max N n)" in allE) by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3631
      ultimately have "\<exists>y\<in>s n. dist y l < e" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3632
    }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3633
    hence "l \<in> s n" using closed_approachable[of "s n" l] assms(1) by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3634
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3635
  thus ?thesis by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3636
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3637
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  3638
text {* Decreasing case does not even need compactness, just completeness. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3639
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3640
lemma decreasing_closed_nest:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3641
  assumes "\<forall>n. closed(s n)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3642
          "\<forall>n. (s n \<noteq> {})"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3643
          "\<forall>m n. m \<le> n --> s n \<subseteq> s m"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3644
          "\<forall>e>0. \<exists>n. \<forall>x \<in> (s n). \<forall> y \<in> (s n). dist x y < e"
44632
076a45f65e12 simplify/generalize some proofs
huffman
parents: 44628
diff changeset
  3645
  shows "\<exists>a::'a::complete_space. \<forall>n::nat. a \<in> s n"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3646
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3647
  have "\<forall>n. \<exists> x. x\<in>s n" using assms(2) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3648
  hence "\<exists>t. \<forall>n. t n \<in> s n" using choice[of "\<lambda> n x. x \<in> s n"] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3649
  then obtain t where t: "\<forall>n. t n \<in> s n" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3650
  { fix e::real assume "e>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3651
    then obtain N where N:"\<forall>x\<in>s N. \<forall>y\<in>s N. dist x y < e" using assms(4) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3652
    { fix m n ::nat assume "N \<le> m \<and> N \<le> n"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3653
      hence "t m \<in> s N" "t n \<in> s N" using assms(3) t unfolding  subset_eq t by blast+
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3654
      hence "dist (t m) (t n) < e" using N by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3655
    }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3656
    hence "\<exists>N. \<forall>m n. N \<le> m \<and> N \<le> n \<longrightarrow> dist (t m) (t n) < e" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3657
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3658
  hence  "Cauchy t" unfolding cauchy_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3659
  then obtain l where l:"(t ---> l) sequentially" using complete_univ unfolding complete_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3660
  { fix n::nat
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3661
    { fix e::real assume "e>0"
44907
93943da0a010 remove redundant lemma Lim_sequentially in favor of lemma LIMSEQ_def
huffman
parents: 44905
diff changeset
  3662
      then obtain N::nat where N:"\<forall>n\<ge>N. dist (t n) l < e" using l[unfolded LIMSEQ_def] by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3663
      have "t (max n N) \<in> s n" using assms(3) unfolding subset_eq apply(erule_tac x=n in allE) apply (erule_tac x="max n N" in allE) using t by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3664
      hence "\<exists>y\<in>s n. dist y l < e" apply(rule_tac x="t (max n N)" in bexI) using N by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3665
    }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3666
    hence "l \<in> s n" using closed_approachable[of "s n" l] assms(1) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3667
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3668
  then show ?thesis by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3669
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3670
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  3671
text {* Strengthen it to the intersection actually being a singleton. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3672
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3673
lemma decreasing_closed_nest_sing:
44632
076a45f65e12 simplify/generalize some proofs
huffman
parents: 44628
diff changeset
  3674
  fixes s :: "nat \<Rightarrow> 'a::complete_space set"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3675
  assumes "\<forall>n. closed(s n)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3676
          "\<forall>n. s n \<noteq> {}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3677
          "\<forall>m n. m \<le> n --> s n \<subseteq> s m"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3678
          "\<forall>e>0. \<exists>n. \<forall>x \<in> (s n). \<forall> y\<in>(s n). dist x y < e"
34104
22758f95e624 re-state lemmas using 'range'
huffman
parents: 33758
diff changeset
  3679
  shows "\<exists>a. \<Inter>(range s) = {a}"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3680
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3681
  obtain a where a:"\<forall>n. a \<in> s n" using decreasing_closed_nest[of s] using assms by auto
34104
22758f95e624 re-state lemmas using 'range'
huffman
parents: 33758
diff changeset
  3682
  { fix b assume b:"b \<in> \<Inter>(range s)"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3683
    { fix e::real assume "e>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3684
      hence "dist a b < e" using assms(4 )using b using a by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3685
    }
36778
739a9379e29b avoid using real-specific versions of generic lemmas
huffman
parents: 36669
diff changeset
  3686
    hence "dist a b = 0" by (metis dist_eq_0_iff dist_nz less_le)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3687
  }
34104
22758f95e624 re-state lemmas using 'range'
huffman
parents: 33758
diff changeset
  3688
  with a have "\<Inter>(range s) = {a}" unfolding image_def by auto
22758f95e624 re-state lemmas using 'range'
huffman
parents: 33758
diff changeset
  3689
  thus ?thesis ..
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3690
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3691
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3692
text{* Cauchy-type criteria for uniform convergence. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3693
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3694
lemma uniformly_convergent_eq_cauchy: fixes s::"nat \<Rightarrow> 'b \<Rightarrow> 'a::heine_borel" shows
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3695
 "(\<exists>l. \<forall>e>0. \<exists>N. \<forall>n x. N \<le> n \<and> P x --> dist(s n x)(l x) < e) \<longleftrightarrow>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3696
  (\<forall>e>0. \<exists>N. \<forall>m n x. N \<le> m \<and> N \<le> n \<and> P x  --> dist (s m x) (s n x) < e)" (is "?lhs = ?rhs")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3697
proof(rule)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3698
  assume ?lhs
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3699
  then obtain l where l:"\<forall>e>0. \<exists>N. \<forall>n x. N \<le> n \<and> P x \<longrightarrow> dist (s n x) (l x) < e" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3700
  { fix e::real assume "e>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3701
    then obtain N::nat where N:"\<forall>n x. N \<le> n \<and> P x \<longrightarrow> dist (s n x) (l x) < e / 2" using l[THEN spec[where x="e/2"]] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3702
    { fix n m::nat and x::"'b" assume "N \<le> m \<and> N \<le> n \<and> P x"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3703
      hence "dist (s m x) (s n x) < e"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3704
        using N[THEN spec[where x=m], THEN spec[where x=x]]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3705
        using N[THEN spec[where x=n], THEN spec[where x=x]]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3706
        using dist_triangle_half_l[of "s m x" "l x" e "s n x"] by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3707
    hence "\<exists>N. \<forall>m n x. N \<le> m \<and> N \<le> n \<and> P x  --> dist (s m x) (s n x) < e"  by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3708
  thus ?rhs by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3709
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3710
  assume ?rhs
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3711
  hence "\<forall>x. P x \<longrightarrow> Cauchy (\<lambda>n. s n x)" unfolding cauchy_def apply auto by (erule_tac x=e in allE)auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3712
  then obtain l where l:"\<forall>x. P x \<longrightarrow> ((\<lambda>n. s n x) ---> l x) sequentially" unfolding convergent_eq_cauchy[THEN sym]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3713
    using choice[of "\<lambda>x l. P x \<longrightarrow> ((\<lambda>n. s n x) ---> l) sequentially"] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3714
  { fix e::real assume "e>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3715
    then obtain N where N:"\<forall>m n x. N \<le> m \<and> N \<le> n \<and> P x \<longrightarrow> dist (s m x) (s n x) < e/2"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3716
      using `?rhs`[THEN spec[where x="e/2"]] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3717
    { fix x assume "P x"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3718
      then obtain M where M:"\<forall>n\<ge>M. dist (s n x) (l x) < e/2"
44907
93943da0a010 remove redundant lemma Lim_sequentially in favor of lemma LIMSEQ_def
huffman
parents: 44905
diff changeset
  3719
        using l[THEN spec[where x=x], unfolded LIMSEQ_def] using `e>0` by(auto elim!: allE[where x="e/2"])
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3720
      fix n::nat assume "n\<ge>N"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3721
      hence "dist(s n x)(l x) < e"  using `P x`and N[THEN spec[where x=n], THEN spec[where x="N+M"], THEN spec[where x=x]]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3722
        using M[THEN spec[where x="N+M"]] and dist_triangle_half_l[of "s n x" "s (N+M) x" e "l x"] by (auto simp add: dist_commute)  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3723
    hence "\<exists>N. \<forall>n x. N \<le> n \<and> P x \<longrightarrow> dist(s n x)(l x) < e" by auto }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3724
  thus ?lhs by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3725
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3726
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3727
lemma uniformly_cauchy_imp_uniformly_convergent:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3728
  fixes s :: "nat \<Rightarrow> 'a \<Rightarrow> 'b::heine_borel"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3729
  assumes "\<forall>e>0.\<exists>N. \<forall>m (n::nat) x. N \<le> m \<and> N \<le> n \<and> P x --> dist(s m x)(s n x) < e"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3730
          "\<forall>x. P x --> (\<forall>e>0. \<exists>N. \<forall>n. N \<le> n --> dist(s n x)(l x) < e)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3731
  shows "\<forall>e>0. \<exists>N. \<forall>n x. N \<le> n \<and> P x --> dist(s n x)(l x) < e"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3732
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3733
  obtain l' where l:"\<forall>e>0. \<exists>N. \<forall>n x. N \<le> n \<and> P x \<longrightarrow> dist (s n x) (l' x) < e"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3734
    using assms(1) unfolding uniformly_convergent_eq_cauchy[THEN sym] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3735
  moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3736
  { fix x assume "P x"
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41969
diff changeset
  3737
    hence "l x = l' x" using tendsto_unique[OF trivial_limit_sequentially, of "\<lambda>n. s n x" "l x" "l' x"]
44907
93943da0a010 remove redundant lemma Lim_sequentially in favor of lemma LIMSEQ_def
huffman
parents: 44905
diff changeset
  3738
      using l and assms(2) unfolding LIMSEQ_def by blast  }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3739
  ultimately show ?thesis by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3740
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3741
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  3742
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  3743
subsection {* Continuity *}
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  3744
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  3745
text {* Define continuity over a net to take in restrictions of the set. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3746
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3747
definition
44081
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44076
diff changeset
  3748
  continuous :: "'a::t2_space filter \<Rightarrow> ('a \<Rightarrow> 'b::topological_space) \<Rightarrow> bool"
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44076
diff changeset
  3749
  where "continuous net f \<longleftrightarrow> (f ---> f(netlimit net)) net"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3750
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3751
lemma continuous_trivial_limit:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3752
 "trivial_limit net ==> continuous net f"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3753
  unfolding continuous_def tendsto_def trivial_limit_eq by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3754
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3755
lemma continuous_within: "continuous (at x within s) f \<longleftrightarrow> (f ---> f(x)) (at x within s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3756
  unfolding continuous_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3757
  unfolding tendsto_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3758
  using netlimit_within[of x s]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3759
  by (cases "trivial_limit (at x within s)") (auto simp add: trivial_limit_eventually)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3760
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3761
lemma continuous_at: "continuous (at x) f \<longleftrightarrow> (f ---> f(x)) (at x)"
45031
9583f2b56f85 add lemmas within_empty and tendsto_bot;
huffman
parents: 44909
diff changeset
  3762
  using continuous_within [of x UNIV f] by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3763
50973
4a2c82644889 generalized diameter from real_normed_vector to metric_space
hoelzl
parents: 50972
diff changeset
  3764
lemma continuous_isCont: "isCont f x = continuous (at x) f"
4a2c82644889 generalized diameter from real_normed_vector to metric_space
hoelzl
parents: 50972
diff changeset
  3765
  unfolding isCont_def LIM_def
4a2c82644889 generalized diameter from real_normed_vector to metric_space
hoelzl
parents: 50972
diff changeset
  3766
  unfolding continuous_at Lim_at unfolding dist_nz by auto
4a2c82644889 generalized diameter from real_normed_vector to metric_space
hoelzl
parents: 50972
diff changeset
  3767
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3768
lemma continuous_at_within:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3769
  assumes "continuous (at x) f"  shows "continuous (at x within s) f"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3770
  using assms unfolding continuous_at continuous_within
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3771
  by (rule Lim_at_within)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3772
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3773
text{* Derive the epsilon-delta forms, which we often use as "definitions" *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3774
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3775
lemma continuous_within_eps_delta:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3776
  "continuous (at x within s) f \<longleftrightarrow> (\<forall>e>0. \<exists>d>0. \<forall>x'\<in> s.  dist x' x < d --> dist (f x') (f x) < e)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3777
  unfolding continuous_within and Lim_within
44584
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  3778
  apply auto unfolding dist_nz[THEN sym] apply(auto del: allE elim!:allE) apply(rule_tac x=d in exI) by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3779
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3780
lemma continuous_at_eps_delta: "continuous (at x) f \<longleftrightarrow>  (\<forall>e>0. \<exists>d>0.
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3781
                           \<forall>x'. dist x' x < d --> dist(f x')(f x) < e)"
45031
9583f2b56f85 add lemmas within_empty and tendsto_bot;
huffman
parents: 44909
diff changeset
  3782
  using continuous_within_eps_delta [of x UNIV f] by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3783
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3784
text{* Versions in terms of open balls. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3785
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3786
lemma continuous_within_ball:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3787
 "continuous (at x within s) f \<longleftrightarrow> (\<forall>e>0. \<exists>d>0.
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3788
                            f ` (ball x d \<inter> s) \<subseteq> ball (f x) e)" (is "?lhs = ?rhs")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3789
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3790
  assume ?lhs
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3791
  { fix e::real assume "e>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3792
    then obtain d where d: "d>0" "\<forall>xa\<in>s. 0 < dist xa x \<and> dist xa x < d \<longrightarrow> dist (f xa) (f x) < e"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3793
      using `?lhs`[unfolded continuous_within Lim_within] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3794
    { fix y assume "y\<in>f ` (ball x d \<inter> s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3795
      hence "y \<in> ball (f x) e" using d(2) unfolding dist_nz[THEN sym]
36362
06475a1547cb fix lots of looping simp calls and other warnings
huffman
parents: 36360
diff changeset
  3796
        apply (auto simp add: dist_commute) apply(erule_tac x=xa in ballE) apply auto using `e>0` by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3797
    }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3798
    hence "\<exists>d>0. f ` (ball x d \<inter> s) \<subseteq> ball (f x) e" using `d>0` unfolding subset_eq ball_def by (auto simp add: dist_commute)  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3799
  thus ?rhs by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3800
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3801
  assume ?rhs thus ?lhs unfolding continuous_within Lim_within ball_def subset_eq
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3802
    apply (auto simp add: dist_commute) apply(erule_tac x=e in allE) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3803
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3804
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3805
lemma continuous_at_ball:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3806
  "continuous (at x) f \<longleftrightarrow> (\<forall>e>0. \<exists>d>0. f ` (ball x d) \<subseteq> ball (f x) e)" (is "?lhs = ?rhs")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3807
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3808
  assume ?lhs thus ?rhs unfolding continuous_at Lim_at subset_eq Ball_def Bex_def image_iff mem_ball
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3809
    apply auto apply(erule_tac x=e in allE) apply auto apply(rule_tac x=d in exI) apply auto apply(erule_tac x=xa in allE) apply (auto simp add: dist_commute dist_nz)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3810
    unfolding dist_nz[THEN sym] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3811
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3812
  assume ?rhs thus ?lhs unfolding continuous_at Lim_at subset_eq Ball_def Bex_def image_iff mem_ball
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3813
    apply auto apply(erule_tac x=e in allE) apply auto apply(rule_tac x=d in exI) apply auto apply(erule_tac x="f xa" in allE) by (auto simp add: dist_commute dist_nz)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3814
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3815
36440
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3816
text{* Define setwise continuity in terms of limits within the set. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3817
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3818
definition
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  3819
  continuous_on ::
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  3820
    "'a set \<Rightarrow> ('a::topological_space \<Rightarrow> 'b::topological_space) \<Rightarrow> bool"
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  3821
where
36440
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3822
  "continuous_on s f \<longleftrightarrow> (\<forall>x\<in>s. (f ---> f x) (at x within s))"
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3823
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3824
lemma continuous_on_topological:
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  3825
  "continuous_on s f \<longleftrightarrow>
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  3826
    (\<forall>x\<in>s. \<forall>B. open B \<longrightarrow> f x \<in> B \<longrightarrow>
36440
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3827
      (\<exists>A. open A \<and> x \<in> A \<and> (\<forall>y\<in>s. y \<in> A \<longrightarrow> f y \<in> B)))"
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3828
unfolding continuous_on_def tendsto_def
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3829
unfolding Limits.eventually_within eventually_at_topological
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3830
by (intro ball_cong [OF refl] all_cong imp_cong ex_cong conj_cong refl) auto
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  3831
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  3832
lemma continuous_on_iff:
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  3833
  "continuous_on s f \<longleftrightarrow>
36440
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3834
    (\<forall>x\<in>s. \<forall>e>0. \<exists>d>0. \<forall>x'\<in>s. dist x' x < d \<longrightarrow> dist (f x') (f x) < e)"
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3835
unfolding continuous_on_def Lim_within
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3836
apply (intro ball_cong [OF refl] all_cong ex_cong)
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3837
apply (rename_tac y, case_tac "y = x", simp)
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3838
apply (simp add: dist_nz)
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  3839
done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3840
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3841
definition
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3842
  uniformly_continuous_on ::
36440
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3843
    "'a set \<Rightarrow> ('a::metric_space \<Rightarrow> 'b::metric_space) \<Rightarrow> bool"
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3844
where
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3845
  "uniformly_continuous_on s f \<longleftrightarrow>
36440
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3846
    (\<forall>e>0. \<exists>d>0. \<forall>x\<in>s. \<forall>x'\<in>s. dist x' x < d \<longrightarrow> dist (f x') (f x) < e)"
35172
579dd5570f96 Added integration to Multivariate-Analysis (upto FTC)
himmelma
parents: 35028
diff changeset
  3847
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3848
text{* Some simple consequential lemmas. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3849
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3850
lemma uniformly_continuous_imp_continuous:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3851
 " uniformly_continuous_on s f ==> continuous_on s f"
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  3852
  unfolding uniformly_continuous_on_def continuous_on_iff by blast
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3853
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3854
lemma continuous_at_imp_continuous_within:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3855
 "continuous (at x) f ==> continuous (at x within s) f"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3856
  unfolding continuous_within continuous_at using Lim_at_within by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3857
36440
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3858
lemma Lim_trivial_limit: "trivial_limit net \<Longrightarrow> (f ---> l) net"
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3859
unfolding tendsto_def by (simp add: trivial_limit_eq)
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3860
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  3861
lemma continuous_at_imp_continuous_on:
36440
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3862
  assumes "\<forall>x\<in>s. continuous (at x) f"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3863
  shows "continuous_on s f"
36440
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3864
unfolding continuous_on_def
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3865
proof
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3866
  fix x assume "x \<in> s"
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3867
  with assms have *: "(f ---> f (netlimit (at x))) (at x)"
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3868
    unfolding continuous_def by simp
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3869
  have "(f ---> f x) (at x)"
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3870
  proof (cases "trivial_limit (at x)")
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3871
    case True thus ?thesis
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3872
      by (rule Lim_trivial_limit)
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3873
  next
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3874
    case False
36667
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  3875
    hence 1: "netlimit (at x) = x"
45031
9583f2b56f85 add lemmas within_empty and tendsto_bot;
huffman
parents: 44909
diff changeset
  3876
      using netlimit_within [of x UNIV] by simp
36440
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3877
    with * show ?thesis by simp
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3878
  qed
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3879
  thus "(f ---> f x) (at x within s)"
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3880
    by (rule Lim_at_within)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3881
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3882
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3883
lemma continuous_on_eq_continuous_within:
36440
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3884
  "continuous_on s f \<longleftrightarrow> (\<forall>x \<in> s. continuous (at x within s) f)"
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3885
unfolding continuous_on_def continuous_def
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3886
apply (rule ball_cong [OF refl])
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3887
apply (case_tac "trivial_limit (at x within s)")
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3888
apply (simp add: Lim_trivial_limit)
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3889
apply (simp add: netlimit_within)
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3890
done
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3891
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3892
lemmas continuous_on = continuous_on_def -- "legacy theorem name"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3893
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3894
lemma continuous_on_eq_continuous_at:
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  3895
  shows "open s ==> (continuous_on s f \<longleftrightarrow> (\<forall>x \<in> s. continuous (at x) f))"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3896
  by (auto simp add: continuous_on continuous_at Lim_within_open)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3897
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3898
lemma continuous_within_subset:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3899
 "continuous (at x within s) f \<Longrightarrow> t \<subseteq> s
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3900
             ==> continuous (at x within t) f"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3901
  unfolding continuous_within by(metis Lim_within_subset)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3902
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3903
lemma continuous_on_subset:
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  3904
  shows "continuous_on s f \<Longrightarrow> t \<subseteq> s ==> continuous_on t f"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3905
  unfolding continuous_on by (metis subset_eq Lim_within_subset)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3906
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3907
lemma continuous_on_interior:
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  3908
  shows "continuous_on s f \<Longrightarrow> x \<in> interior s \<Longrightarrow> continuous (at x) f"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  3909
  by (erule interiorE, drule (1) continuous_on_subset,
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  3910
    simp add: continuous_on_eq_continuous_at)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3911
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3912
lemma continuous_on_eq:
36440
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3913
  "(\<forall>x \<in> s. f x = g x) \<Longrightarrow> continuous_on s f \<Longrightarrow> continuous_on s g"
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3914
  unfolding continuous_on_def tendsto_def Limits.eventually_within
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3915
  by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3916
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  3917
text {* Characterization of various kinds of continuity in terms of sequences. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3918
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3919
lemma continuous_within_sequentially:
44533
7abe4a59f75d generalize and simplify proof of continuous_within_sequentially
huffman
parents: 44531
diff changeset
  3920
  fixes f :: "'a::metric_space \<Rightarrow> 'b::topological_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3921
  shows "continuous (at a within s) f \<longleftrightarrow>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3922
                (\<forall>x. (\<forall>n::nat. x n \<in> s) \<and> (x ---> a) sequentially
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3923
                     --> ((f o x) ---> f a) sequentially)" (is "?lhs = ?rhs")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3924
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3925
  assume ?lhs
44533
7abe4a59f75d generalize and simplify proof of continuous_within_sequentially
huffman
parents: 44531
diff changeset
  3926
  { fix x::"nat \<Rightarrow> 'a" assume x:"\<forall>n. x n \<in> s" "\<forall>e>0. eventually (\<lambda>n. dist (x n) a < e) sequentially"
7abe4a59f75d generalize and simplify proof of continuous_within_sequentially
huffman
parents: 44531
diff changeset
  3927
    fix T::"'b set" assume "open T" and "f a \<in> T"
7abe4a59f75d generalize and simplify proof of continuous_within_sequentially
huffman
parents: 44531
diff changeset
  3928
    with `?lhs` obtain d where "d>0" and d:"\<forall>x\<in>s. 0 < dist x a \<and> dist x a < d \<longrightarrow> f x \<in> T"
7abe4a59f75d generalize and simplify proof of continuous_within_sequentially
huffman
parents: 44531
diff changeset
  3929
      unfolding continuous_within tendsto_def eventually_within by auto
7abe4a59f75d generalize and simplify proof of continuous_within_sequentially
huffman
parents: 44531
diff changeset
  3930
    have "eventually (\<lambda>n. dist (x n) a < d) sequentially"
7abe4a59f75d generalize and simplify proof of continuous_within_sequentially
huffman
parents: 44531
diff changeset
  3931
      using x(2) `d>0` by simp
7abe4a59f75d generalize and simplify proof of continuous_within_sequentially
huffman
parents: 44531
diff changeset
  3932
    hence "eventually (\<lambda>n. (f \<circ> x) n \<in> T) sequentially"
46887
cb891d9a23c1 use eventually_elim method
noschinl
parents: 45776
diff changeset
  3933
    proof eventually_elim
cb891d9a23c1 use eventually_elim method
noschinl
parents: 45776
diff changeset
  3934
      case (elim n) thus ?case
44533
7abe4a59f75d generalize and simplify proof of continuous_within_sequentially
huffman
parents: 44531
diff changeset
  3935
        using d x(1) `f a \<in> T` unfolding dist_nz[THEN sym] by auto
7abe4a59f75d generalize and simplify proof of continuous_within_sequentially
huffman
parents: 44531
diff changeset
  3936
    qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3937
  }
44533
7abe4a59f75d generalize and simplify proof of continuous_within_sequentially
huffman
parents: 44531
diff changeset
  3938
  thus ?rhs unfolding tendsto_iff unfolding tendsto_def by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3939
next
44533
7abe4a59f75d generalize and simplify proof of continuous_within_sequentially
huffman
parents: 44531
diff changeset
  3940
  assume ?rhs thus ?lhs
7abe4a59f75d generalize and simplify proof of continuous_within_sequentially
huffman
parents: 44531
diff changeset
  3941
    unfolding continuous_within tendsto_def [where l="f a"]
7abe4a59f75d generalize and simplify proof of continuous_within_sequentially
huffman
parents: 44531
diff changeset
  3942
    by (simp add: sequentially_imp_eventually_within)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3943
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3944
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3945
lemma continuous_at_sequentially:
44533
7abe4a59f75d generalize and simplify proof of continuous_within_sequentially
huffman
parents: 44531
diff changeset
  3946
  fixes f :: "'a::metric_space \<Rightarrow> 'b::topological_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3947
  shows "continuous (at a) f \<longleftrightarrow> (\<forall>x. (x ---> a) sequentially
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3948
                  --> ((f o x) ---> f a) sequentially)"
45031
9583f2b56f85 add lemmas within_empty and tendsto_bot;
huffman
parents: 44909
diff changeset
  3949
  using continuous_within_sequentially[of a UNIV f] by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3950
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3951
lemma continuous_on_sequentially:
44533
7abe4a59f75d generalize and simplify proof of continuous_within_sequentially
huffman
parents: 44531
diff changeset
  3952
  fixes f :: "'a::metric_space \<Rightarrow> 'b::topological_space"
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  3953
  shows "continuous_on s f \<longleftrightarrow>
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  3954
    (\<forall>x. \<forall>a \<in> s. (\<forall>n. x(n) \<in> s) \<and> (x ---> a) sequentially
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3955
                    --> ((f o x) ---> f(a)) sequentially)" (is "?lhs = ?rhs")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3956
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3957
  assume ?rhs thus ?lhs using continuous_within_sequentially[of _ s f] unfolding continuous_on_eq_continuous_within by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3958
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3959
  assume ?lhs thus ?rhs unfolding continuous_on_eq_continuous_within using continuous_within_sequentially[of _ s f] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3960
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3961
44648
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  3962
lemma uniformly_continuous_on_sequentially:
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  3963
  "uniformly_continuous_on s f \<longleftrightarrow> (\<forall>x y. (\<forall>n. x n \<in> s) \<and> (\<forall>n. y n \<in> s) \<and>
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  3964
                    ((\<lambda>n. dist (x n) (y n)) ---> 0) sequentially
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  3965
                    \<longrightarrow> ((\<lambda>n. dist (f(x n)) (f(y n))) ---> 0) sequentially)" (is "?lhs = ?rhs")
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3966
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3967
  assume ?lhs
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  3968
  { fix x y assume x:"\<forall>n. x n \<in> s" and y:"\<forall>n. y n \<in> s" and xy:"((\<lambda>n. dist (x n) (y n)) ---> 0) sequentially"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3969
    { fix e::real assume "e>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3970
      then obtain d where "d>0" and d:"\<forall>x\<in>s. \<forall>x'\<in>s. dist x' x < d \<longrightarrow> dist (f x') (f x) < e"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3971
        using `?lhs`[unfolded uniformly_continuous_on_def, THEN spec[where x=e]] by auto
44907
93943da0a010 remove redundant lemma Lim_sequentially in favor of lemma LIMSEQ_def
huffman
parents: 44905
diff changeset
  3972
      obtain N where N:"\<forall>n\<ge>N. dist (x n) (y n) < d" using xy[unfolded LIMSEQ_def dist_norm] and `d>0` by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3973
      { fix n assume "n\<ge>N"
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  3974
        hence "dist (f (x n)) (f (y n)) < e"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3975
          using N[THEN spec[where x=n]] using d[THEN bspec[where x="x n"], THEN bspec[where x="y n"]] using x and y
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  3976
          unfolding dist_commute by simp  }
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  3977
      hence "\<exists>N. \<forall>n\<ge>N. dist (f (x n)) (f (y n)) < e"  by auto  }
44907
93943da0a010 remove redundant lemma Lim_sequentially in favor of lemma LIMSEQ_def
huffman
parents: 44905
diff changeset
  3978
    hence "((\<lambda>n. dist (f(x n)) (f(y n))) ---> 0) sequentially" unfolding LIMSEQ_def and dist_real_def by auto  }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3979
  thus ?rhs by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3980
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3981
  assume ?rhs
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3982
  { assume "\<not> ?lhs"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3983
    then obtain e where "e>0" "\<forall>d>0. \<exists>x\<in>s. \<exists>x'\<in>s. dist x' x < d \<and> \<not> dist (f x') (f x) < e" unfolding uniformly_continuous_on_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3984
    then obtain fa where fa:"\<forall>x.  0 < x \<longrightarrow> fst (fa x) \<in> s \<and> snd (fa x) \<in> s \<and> dist (fst (fa x)) (snd (fa x)) < x \<and> \<not> dist (f (fst (fa x))) (f (snd (fa x))) < e"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3985
      using choice[of "\<lambda>d x. d>0 \<longrightarrow> fst x \<in> s \<and> snd x \<in> s \<and> dist (snd x) (fst x) < d \<and> \<not> dist (f (snd x)) (f (fst x)) < e"] unfolding Bex_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3986
      by (auto simp add: dist_commute)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3987
    def x \<equiv> "\<lambda>n::nat. fst (fa (inverse (real n + 1)))"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3988
    def y \<equiv> "\<lambda>n::nat. snd (fa (inverse (real n + 1)))"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3989
    have xyn:"\<forall>n. x n \<in> s \<and> y n \<in> s" and xy0:"\<forall>n. dist (x n) (y n) < inverse (real n + 1)" and fxy:"\<forall>n. \<not> dist (f (x n)) (f (y n)) < e"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3990
      unfolding x_def and y_def using fa by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3991
    { fix e::real assume "e>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3992
      then obtain N::nat where "N \<noteq> 0" and N:"0 < inverse (real N) \<and> inverse (real N) < e" unfolding real_arch_inv[of e]   by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3993
      { fix n::nat assume "n\<ge>N"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3994
        hence "inverse (real n + 1) < inverse (real N)" using real_of_nat_ge_zero and `N\<noteq>0` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3995
        also have "\<dots> < e" using N by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3996
        finally have "inverse (real n + 1) < e" by auto
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  3997
        hence "dist (x n) (y n) < e" using xy0[THEN spec[where x=n]] by auto  }
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  3998
      hence "\<exists>N. \<forall>n\<ge>N. dist (x n) (y n) < e" by auto  }
44907
93943da0a010 remove redundant lemma Lim_sequentially in favor of lemma LIMSEQ_def
huffman
parents: 44905
diff changeset
  3999
    hence "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. dist (f (x n)) (f (y n)) < e" using `?rhs`[THEN spec[where x=x], THEN spec[where x=y]] and xyn unfolding LIMSEQ_def dist_real_def by auto
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4000
    hence False using fxy and `e>0` by auto  }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4001
  thus ?lhs unfolding uniformly_continuous_on_def by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4002
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4003
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4004
text{* The usual transformation theorems. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4005
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4006
lemma continuous_transform_within:
36667
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  4007
  fixes f g :: "'a::metric_space \<Rightarrow> 'b::topological_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4008
  assumes "0 < d" "x \<in> s" "\<forall>x' \<in> s. dist x' x < d --> f x' = g x'"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4009
          "continuous (at x within s) f"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4010
  shows "continuous (at x within s) g"
36667
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  4011
unfolding continuous_within
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  4012
proof (rule Lim_transform_within)
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  4013
  show "0 < d" by fact
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  4014
  show "\<forall>x'\<in>s. 0 < dist x' x \<and> dist x' x < d \<longrightarrow> f x' = g x'"
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  4015
    using assms(3) by auto
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  4016
  have "f x = g x"
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  4017
    using assms(1,2,3) by auto
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  4018
  thus "(f ---> g x) (at x within s)"
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  4019
    using assms(4) unfolding continuous_within by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4020
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4021
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4022
lemma continuous_transform_at:
36667
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  4023
  fixes f g :: "'a::metric_space \<Rightarrow> 'b::topological_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4024
  assumes "0 < d" "\<forall>x'. dist x' x < d --> f x' = g x'"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4025
          "continuous (at x) f"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4026
  shows "continuous (at x) g"
45031
9583f2b56f85 add lemmas within_empty and tendsto_bot;
huffman
parents: 44909
diff changeset
  4027
  using continuous_transform_within [of d x UNIV f g] assms by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4028
44648
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4029
subsubsection {* Structural rules for pointwise continuity *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4030
44647
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4031
lemma continuous_within_id: "continuous (at a within s) (\<lambda>x. x)"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4032
  unfolding continuous_within by (rule tendsto_ident_at_within)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4033
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4034
lemma continuous_at_id: "continuous (at a) (\<lambda>x. x)"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4035
  unfolding continuous_at by (rule tendsto_ident_at)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4036
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4037
lemma continuous_const: "continuous F (\<lambda>x. c)"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4038
  unfolding continuous_def by (rule tendsto_const)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4039
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4040
lemma continuous_dist:
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4041
  assumes "continuous F f" and "continuous F g"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4042
  shows "continuous F (\<lambda>x. dist (f x) (g x))"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4043
  using assms unfolding continuous_def by (rule tendsto_dist)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4044
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  4045
lemma continuous_infdist:
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  4046
  assumes "continuous F f"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  4047
  shows "continuous F (\<lambda>x. infdist (f x) A)"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  4048
  using assms unfolding continuous_def by (rule tendsto_infdist)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  4049
44647
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4050
lemma continuous_norm:
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4051
  shows "continuous F f \<Longrightarrow> continuous F (\<lambda>x. norm (f x))"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4052
  unfolding continuous_def by (rule tendsto_norm)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4053
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4054
lemma continuous_infnorm:
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4055
  shows "continuous F f \<Longrightarrow> continuous F (\<lambda>x. infnorm (f x))"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4056
  unfolding continuous_def by (rule tendsto_infnorm)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4057
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4058
lemma continuous_add:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4059
  fixes f g :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
44647
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4060
  shows "\<lbrakk>continuous F f; continuous F g\<rbrakk> \<Longrightarrow> continuous F (\<lambda>x. f x + g x)"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4061
  unfolding continuous_def by (rule tendsto_add)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4062
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4063
lemma continuous_minus:
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4064
  fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4065
  shows "continuous F f \<Longrightarrow> continuous F (\<lambda>x. - f x)"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4066
  unfolding continuous_def by (rule tendsto_minus)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4067
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4068
lemma continuous_diff:
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4069
  fixes f g :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
44647
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4070
  shows "\<lbrakk>continuous F f; continuous F g\<rbrakk> \<Longrightarrow> continuous F (\<lambda>x. f x - g x)"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4071
  unfolding continuous_def by (rule tendsto_diff)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4072
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4073
lemma continuous_scaleR:
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4074
  fixes g :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4075
  shows "\<lbrakk>continuous F f; continuous F g\<rbrakk> \<Longrightarrow> continuous F (\<lambda>x. f x *\<^sub>R g x)"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4076
  unfolding continuous_def by (rule tendsto_scaleR)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4077
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4078
lemma continuous_mult:
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4079
  fixes f g :: "'a::t2_space \<Rightarrow> 'b::real_normed_algebra"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4080
  shows "\<lbrakk>continuous F f; continuous F g\<rbrakk> \<Longrightarrow> continuous F (\<lambda>x. f x * g x)"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4081
  unfolding continuous_def by (rule tendsto_mult)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4082
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4083
lemma continuous_inner:
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4084
  assumes "continuous F f" and "continuous F g"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4085
  shows "continuous F (\<lambda>x. inner (f x) (g x))"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4086
  using assms unfolding continuous_def by (rule tendsto_inner)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4087
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4088
lemma continuous_inverse:
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4089
  fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_div_algebra"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4090
  assumes "continuous F f" and "f (netlimit F) \<noteq> 0"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4091
  shows "continuous F (\<lambda>x. inverse (f x))"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4092
  using assms unfolding continuous_def by (rule tendsto_inverse)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4093
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4094
lemma continuous_at_within_inverse:
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4095
  fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_div_algebra"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4096
  assumes "continuous (at a within s) f" and "f a \<noteq> 0"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4097
  shows "continuous (at a within s) (\<lambda>x. inverse (f x))"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4098
  using assms unfolding continuous_within by (rule tendsto_inverse)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4099
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4100
lemma continuous_at_inverse:
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4101
  fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_div_algebra"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4102
  assumes "continuous (at a) f" and "f a \<noteq> 0"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4103
  shows "continuous (at a) (\<lambda>x. inverse (f x))"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4104
  using assms unfolding continuous_at by (rule tendsto_inverse)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4105
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4106
lemmas continuous_intros = continuous_at_id continuous_within_id
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4107
  continuous_const continuous_dist continuous_norm continuous_infnorm
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  4108
  continuous_add continuous_minus continuous_diff continuous_scaleR continuous_mult
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  4109
  continuous_inner continuous_at_inverse continuous_at_within_inverse
34964
4e8be3c04d37 Replaced vec1 and dest_vec1 by abbreviation.
hoelzl
parents: 34291
diff changeset
  4110
44648
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4111
subsubsection {* Structural rules for setwise continuity *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4112
44647
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4113
lemma continuous_on_id: "continuous_on s (\<lambda>x. x)"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4114
  unfolding continuous_on_def by (fast intro: tendsto_ident_at_within)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4115
44531
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4116
lemma continuous_on_const: "continuous_on s (\<lambda>x. c)"
44125
230a8665c919 mark some redundant theorems as legacy
huffman
parents: 44122
diff changeset
  4117
  unfolding continuous_on_def by (auto intro: tendsto_intros)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4118
44647
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4119
lemma continuous_on_norm:
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4120
  shows "continuous_on s f \<Longrightarrow> continuous_on s (\<lambda>x. norm (f x))"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4121
  unfolding continuous_on_def by (fast intro: tendsto_norm)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4122
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4123
lemma continuous_on_infnorm:
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4124
  shows "continuous_on s f \<Longrightarrow> continuous_on s (\<lambda>x. infnorm (f x))"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4125
  unfolding continuous_on by (fast intro: tendsto_infnorm)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4126
44531
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4127
lemma continuous_on_minus:
36440
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  4128
  fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4129
  shows "continuous_on s f \<Longrightarrow> continuous_on s (\<lambda>x. - f x)"
36440
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  4130
  unfolding continuous_on_def by (auto intro: tendsto_intros)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4131
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4132
lemma continuous_on_add:
36440
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  4133
  fixes f g :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4134
  shows "continuous_on s f \<Longrightarrow> continuous_on s g
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4135
           \<Longrightarrow> continuous_on s (\<lambda>x. f x + g x)"
36440
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  4136
  unfolding continuous_on_def by (auto intro: tendsto_intros)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4137
44531
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4138
lemma continuous_on_diff:
36440
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  4139
  fixes f g :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4140
  shows "continuous_on s f \<Longrightarrow> continuous_on s g
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4141
           \<Longrightarrow> continuous_on s (\<lambda>x. f x - g x)"
36440
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  4142
  unfolding continuous_on_def by (auto intro: tendsto_intros)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4143
44531
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4144
lemma (in bounded_linear) continuous_on:
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4145
  "continuous_on s g \<Longrightarrow> continuous_on s (\<lambda>x. f (g x))"
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4146
  unfolding continuous_on_def by (fast intro: tendsto)
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4147
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4148
lemma (in bounded_bilinear) continuous_on:
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4149
  "\<lbrakk>continuous_on s f; continuous_on s g\<rbrakk> \<Longrightarrow> continuous_on s (\<lambda>x. f x ** g x)"
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4150
  unfolding continuous_on_def by (fast intro: tendsto)
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4151
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4152
lemma continuous_on_scaleR:
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4153
  fixes g :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector"
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4154
  assumes "continuous_on s f" and "continuous_on s g"
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4155
  shows "continuous_on s (\<lambda>x. f x *\<^sub>R g x)"
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4156
  using bounded_bilinear_scaleR assms
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4157
  by (rule bounded_bilinear.continuous_on)
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4158
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4159
lemma continuous_on_mult:
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4160
  fixes g :: "'a::topological_space \<Rightarrow> 'b::real_normed_algebra"
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4161
  assumes "continuous_on s f" and "continuous_on s g"
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4162
  shows "continuous_on s (\<lambda>x. f x * g x)"
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4163
  using bounded_bilinear_mult assms
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4164
  by (rule bounded_bilinear.continuous_on)
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4165
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4166
lemma continuous_on_inner:
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4167
  fixes g :: "'a::topological_space \<Rightarrow> 'b::real_inner"
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4168
  assumes "continuous_on s f" and "continuous_on s g"
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4169
  shows "continuous_on s (\<lambda>x. inner (f x) (g x))"
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4170
  using bounded_bilinear_inner assms
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4171
  by (rule bounded_bilinear.continuous_on)
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4172
44647
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4173
lemma continuous_on_inverse:
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4174
  fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_div_algebra"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4175
  assumes "continuous_on s f" and "\<forall>x\<in>s. f x \<noteq> 0"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4176
  shows "continuous_on s (\<lambda>x. inverse (f x))"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4177
  using assms unfolding continuous_on by (fast intro: tendsto_inverse)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4178
44648
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4179
subsubsection {* Structural rules for uniform continuity *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4180
44647
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4181
lemma uniformly_continuous_on_id:
44648
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4182
  shows "uniformly_continuous_on s (\<lambda>x. x)"
44647
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4183
  unfolding uniformly_continuous_on_def by auto
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4184
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4185
lemma uniformly_continuous_on_const:
44648
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4186
  shows "uniformly_continuous_on s (\<lambda>x. c)"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4187
  unfolding uniformly_continuous_on_def by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4188
44648
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4189
lemma uniformly_continuous_on_dist:
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4190
  fixes f g :: "'a::metric_space \<Rightarrow> 'b::metric_space"
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4191
  assumes "uniformly_continuous_on s f"
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4192
  assumes "uniformly_continuous_on s g"
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4193
  shows "uniformly_continuous_on s (\<lambda>x. dist (f x) (g x))"
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4194
proof -
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4195
  { fix a b c d :: 'b have "\<bar>dist a b - dist c d\<bar> \<le> dist a c + dist b d"
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4196
      using dist_triangle2 [of a b c] dist_triangle2 [of b c d]
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4197
      using dist_triangle3 [of c d a] dist_triangle [of a d b]
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4198
      by arith
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4199
  } note le = this
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4200
  { fix x y
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4201
    assume f: "(\<lambda>n. dist (f (x n)) (f (y n))) ----> 0"
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4202
    assume g: "(\<lambda>n. dist (g (x n)) (g (y n))) ----> 0"
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4203
    have "(\<lambda>n. \<bar>dist (f (x n)) (g (x n)) - dist (f (y n)) (g (y n))\<bar>) ----> 0"
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4204
      by (rule Lim_transform_bound [OF _ tendsto_add_zero [OF f g]],
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4205
        simp add: le)
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4206
  }
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4207
  thus ?thesis using assms unfolding uniformly_continuous_on_sequentially
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4208
    unfolding dist_real_def by simp
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4209
qed
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4210
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4211
lemma uniformly_continuous_on_norm:
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4212
  assumes "uniformly_continuous_on s f"
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4213
  shows "uniformly_continuous_on s (\<lambda>x. norm (f x))"
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4214
  unfolding norm_conv_dist using assms
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4215
  by (intro uniformly_continuous_on_dist uniformly_continuous_on_const)
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4216
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4217
lemma (in bounded_linear) uniformly_continuous_on:
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4218
  assumes "uniformly_continuous_on s g"
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4219
  shows "uniformly_continuous_on s (\<lambda>x. f (g x))"
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4220
  using assms unfolding uniformly_continuous_on_sequentially
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4221
  unfolding dist_norm tendsto_norm_zero_iff diff[symmetric]
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4222
  by (auto intro: tendsto_zero)
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4223
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4224
lemma uniformly_continuous_on_cmul:
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4225
  fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4226
  assumes "uniformly_continuous_on s f"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4227
  shows "uniformly_continuous_on s (\<lambda>x. c *\<^sub>R f(x))"
44648
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4228
  using bounded_linear_scaleR_right assms
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4229
  by (rule bounded_linear.uniformly_continuous_on)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4230
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4231
lemma dist_minus:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4232
  fixes x y :: "'a::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4233
  shows "dist (- x) (- y) = dist x y"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4234
  unfolding dist_norm minus_diff_minus norm_minus_cancel ..
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4235
44648
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4236
lemma uniformly_continuous_on_minus:
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4237
  fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"
44648
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4238
  shows "uniformly_continuous_on s f \<Longrightarrow> uniformly_continuous_on s (\<lambda>x. - f x)"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4239
  unfolding uniformly_continuous_on_def dist_minus .
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4240
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4241
lemma uniformly_continuous_on_add:
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4242
  fixes f g :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"
44648
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4243
  assumes "uniformly_continuous_on s f"
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4244
  assumes "uniformly_continuous_on s g"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4245
  shows "uniformly_continuous_on s (\<lambda>x. f x + g x)"
44648
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4246
  using assms unfolding uniformly_continuous_on_sequentially
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4247
  unfolding dist_norm tendsto_norm_zero_iff add_diff_add
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4248
  by (auto intro: tendsto_add_zero)
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4249
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4250
lemma uniformly_continuous_on_diff:
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4251
  fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"
44648
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4252
  assumes "uniformly_continuous_on s f" and "uniformly_continuous_on s g"
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4253
  shows "uniformly_continuous_on s (\<lambda>x. f x - g x)"
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4254
  unfolding ab_diff_minus using assms
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4255
  by (intro uniformly_continuous_on_add uniformly_continuous_on_minus)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4256
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4257
text{* Continuity of all kinds is preserved under composition. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4258
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4259
lemma continuous_within_topological:
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4260
  "continuous (at x within s) f \<longleftrightarrow>
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4261
    (\<forall>B. open B \<longrightarrow> f x \<in> B \<longrightarrow>
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4262
      (\<exists>A. open A \<and> x \<in> A \<and> (\<forall>y\<in>s. y \<in> A \<longrightarrow> f y \<in> B)))"
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4263
unfolding continuous_within
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4264
unfolding tendsto_def Limits.eventually_within eventually_at_topological
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4265
by (intro ball_cong [OF refl] all_cong imp_cong ex_cong conj_cong refl) auto
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4266
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4267
lemma continuous_within_compose:
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4268
  assumes "continuous (at x within s) f"
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4269
  assumes "continuous (at (f x) within f ` s) g"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4270
  shows "continuous (at x within s) (g o f)"
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4271
using assms unfolding continuous_within_topological by simp metis
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4272
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4273
lemma continuous_at_compose:
45031
9583f2b56f85 add lemmas within_empty and tendsto_bot;
huffman
parents: 44909
diff changeset
  4274
  assumes "continuous (at x) f" and "continuous (at (f x)) g"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4275
  shows "continuous (at x) (g o f)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4276
proof-
45031
9583f2b56f85 add lemmas within_empty and tendsto_bot;
huffman
parents: 44909
diff changeset
  4277
  have "continuous (at (f x) within range f) g" using assms(2)
9583f2b56f85 add lemmas within_empty and tendsto_bot;
huffman
parents: 44909
diff changeset
  4278
    using continuous_within_subset[of "f x" UNIV g "range f"] by simp
9583f2b56f85 add lemmas within_empty and tendsto_bot;
huffman
parents: 44909
diff changeset
  4279
  thus ?thesis using assms(1)
9583f2b56f85 add lemmas within_empty and tendsto_bot;
huffman
parents: 44909
diff changeset
  4280
    using continuous_within_compose[of x UNIV f g] by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4281
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4282
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4283
lemma continuous_on_compose:
36440
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  4284
  "continuous_on s f \<Longrightarrow> continuous_on (f ` s) g \<Longrightarrow> continuous_on s (g o f)"
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  4285
  unfolding continuous_on_topological by simp metis
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4286
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4287
lemma uniformly_continuous_on_compose:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4288
  assumes "uniformly_continuous_on s f"  "uniformly_continuous_on (f ` s) g"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4289
  shows "uniformly_continuous_on s (g o f)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4290
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4291
  { fix e::real assume "e>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4292
    then obtain d where "d>0" and d:"\<forall>x\<in>f ` s. \<forall>x'\<in>f ` s. dist x' x < d \<longrightarrow> dist (g x') (g x) < e" using assms(2) unfolding uniformly_continuous_on_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4293
    obtain d' where "d'>0" "\<forall>x\<in>s. \<forall>x'\<in>s. dist x' x < d' \<longrightarrow> dist (f x') (f x) < d" using `d>0` using assms(1) unfolding uniformly_continuous_on_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4294
    hence "\<exists>d>0. \<forall>x\<in>s. \<forall>x'\<in>s. dist x' x < d \<longrightarrow> dist ((g \<circ> f) x') ((g \<circ> f) x) < e" using `d>0` using d by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4295
  thus ?thesis using assms unfolding uniformly_continuous_on_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4296
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4297
44647
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4298
lemmas continuous_on_intros = continuous_on_id continuous_on_const
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4299
  continuous_on_compose continuous_on_norm continuous_on_infnorm
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4300
  continuous_on_add continuous_on_minus continuous_on_diff
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4301
  continuous_on_scaleR continuous_on_mult continuous_on_inverse
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  4302
  continuous_on_inner
44648
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4303
  uniformly_continuous_on_id uniformly_continuous_on_const
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4304
  uniformly_continuous_on_dist uniformly_continuous_on_norm
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4305
  uniformly_continuous_on_compose uniformly_continuous_on_add
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4306
  uniformly_continuous_on_minus uniformly_continuous_on_diff
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4307
  uniformly_continuous_on_cmul
44647
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4308
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4309
text{* Continuity in terms of open preimages. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4310
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4311
lemma continuous_at_open:
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4312
  shows "continuous (at x) f \<longleftrightarrow> (\<forall>t. open t \<and> f x \<in> t --> (\<exists>s. open s \<and> x \<in> s \<and> (\<forall>x' \<in> s. (f x') \<in> t)))"
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4313
unfolding continuous_within_topological [of x UNIV f, unfolded within_UNIV]
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4314
unfolding imp_conjL by (intro all_cong imp_cong ex_cong conj_cong refl) auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4315
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4316
lemma continuous_on_open:
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4317
  shows "continuous_on s f \<longleftrightarrow>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4318
        (\<forall>t. openin (subtopology euclidean (f ` s)) t
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4319
            --> openin (subtopology euclidean s) {x \<in> s. f x \<in> t})" (is "?lhs = ?rhs")
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4320
proof (safe)
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4321
  fix t :: "'b set"
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4322
  assume 1: "continuous_on s f"
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4323
  assume 2: "openin (subtopology euclidean (f ` s)) t"
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4324
  from 2 obtain B where B: "open B" and t: "t = f ` s \<inter> B"
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4325
    unfolding openin_open by auto
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4326
  def U == "\<Union>{A. open A \<and> (\<forall>x\<in>s. x \<in> A \<longrightarrow> f x \<in> B)}"
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4327
  have "open U" unfolding U_def by (simp add: open_Union)
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4328
  moreover have "\<forall>x\<in>s. x \<in> U \<longleftrightarrow> f x \<in> t"
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4329
  proof (intro ballI iffI)
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4330
    fix x assume "x \<in> s" and "x \<in> U" thus "f x \<in> t"
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4331
      unfolding U_def t by auto
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4332
  next
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4333
    fix x assume "x \<in> s" and "f x \<in> t"
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4334
    hence "x \<in> s" and "f x \<in> B"
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4335
      unfolding t by auto
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4336
    with 1 B obtain A where "open A" "x \<in> A" "\<forall>y\<in>s. y \<in> A \<longrightarrow> f y \<in> B"
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4337
      unfolding t continuous_on_topological by metis
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4338
    then show "x \<in> U"
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4339
      unfolding U_def by auto
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4340
  qed
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4341
  ultimately have "open U \<and> {x \<in> s. f x \<in> t} = s \<inter> U" by auto
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4342
  then show "openin (subtopology euclidean s) {x \<in> s. f x \<in> t}"
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4343
    unfolding openin_open by fast
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4344
next
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4345
  assume "?rhs" show "continuous_on s f"
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4346
  unfolding continuous_on_topological
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4347
  proof (clarify)
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4348
    fix x and B assume "x \<in> s" and "open B" and "f x \<in> B"
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4349
    have "openin (subtopology euclidean (f ` s)) (f ` s \<inter> B)"
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4350
      unfolding openin_open using `open B` by auto
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4351
    then have "openin (subtopology euclidean s) {x \<in> s. f x \<in> f ` s \<inter> B}"
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4352
      using `?rhs` by fast
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4353
    then show "\<exists>A. open A \<and> x \<in> A \<and> (\<forall>y\<in>s. y \<in> A \<longrightarrow> f y \<in> B)"
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4354
      unfolding openin_open using `x \<in> s` and `f x \<in> B` by auto
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4355
  qed
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4356
qed
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4357
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4358
text {* Similarly in terms of closed sets. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4359
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4360
lemma continuous_on_closed:
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  4361
  shows "continuous_on s f \<longleftrightarrow>  (\<forall>t. closedin (subtopology euclidean (f ` s)) t  --> closedin (subtopology euclidean s) {x \<in> s. f x \<in> t})" (is "?lhs = ?rhs")
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4362
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4363
  assume ?lhs
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4364
  { fix t
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4365
    have *:"s - {x \<in> s. f x \<in> f ` s - t} = {x \<in> s. f x \<in> t}" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4366
    have **:"f ` s - (f ` s - (f ` s - t)) = f ` s - t" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4367
    assume as:"closedin (subtopology euclidean (f ` s)) t"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4368
    hence "closedin (subtopology euclidean (f ` s)) (f ` s - (f ` s - t))" unfolding closedin_def topspace_euclidean_subtopology unfolding ** by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4369
    hence "closedin (subtopology euclidean s) {x \<in> s. f x \<in> t}" using `?lhs`[unfolded continuous_on_open, THEN spec[where x="(f ` s) - t"]]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4370
      unfolding openin_closedin_eq topspace_euclidean_subtopology unfolding * by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4371
  thus ?rhs by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4372
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4373
  assume ?rhs
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4374
  { fix t
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4375
    have *:"s - {x \<in> s. f x \<in> f ` s - t} = {x \<in> s. f x \<in> t}" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4376
    assume as:"openin (subtopology euclidean (f ` s)) t"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4377
    hence "openin (subtopology euclidean s) {x \<in> s. f x \<in> t}" using `?rhs`[THEN spec[where x="(f ` s) - t"]]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4378
      unfolding openin_closedin_eq topspace_euclidean_subtopology *[THEN sym] closedin_subtopology by auto }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4379
  thus ?lhs unfolding continuous_on_open by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4380
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4381
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  4382
text {* Half-global and completely global cases. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4383
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4384
lemma continuous_open_in_preimage:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4385
  assumes "continuous_on s f"  "open t"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4386
  shows "openin (subtopology euclidean s) {x \<in> s. f x \<in> t}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4387
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4388
  have *:"\<forall>x. x \<in> s \<and> f x \<in> t \<longleftrightarrow> x \<in> s \<and> f x \<in> (t \<inter> f ` s)" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4389
  have "openin (subtopology euclidean (f ` s)) (t \<inter> f ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4390
    using openin_open_Int[of t "f ` s", OF assms(2)] unfolding openin_open by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4391
  thus ?thesis using assms(1)[unfolded continuous_on_open, THEN spec[where x="t \<inter> f ` s"]] using * by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4392
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4393
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4394
lemma continuous_closed_in_preimage:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4395
  assumes "continuous_on s f"  "closed t"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4396
  shows "closedin (subtopology euclidean s) {x \<in> s. f x \<in> t}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4397
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4398
  have *:"\<forall>x. x \<in> s \<and> f x \<in> t \<longleftrightarrow> x \<in> s \<and> f x \<in> (t \<inter> f ` s)" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4399
  have "closedin (subtopology euclidean (f ` s)) (t \<inter> f ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4400
    using closedin_closed_Int[of t "f ` s", OF assms(2)] unfolding Int_commute by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4401
  thus ?thesis
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4402
    using assms(1)[unfolded continuous_on_closed, THEN spec[where x="t \<inter> f ` s"]] using * by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4403
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4404
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4405
lemma continuous_open_preimage:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4406
  assumes "continuous_on s f" "open s" "open t"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4407
  shows "open {x \<in> s. f x \<in> t}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4408
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4409
  obtain T where T: "open T" "{x \<in> s. f x \<in> t} = s \<inter> T"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4410
    using continuous_open_in_preimage[OF assms(1,3)] unfolding openin_open by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4411
  thus ?thesis using open_Int[of s T, OF assms(2)] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4412
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4413
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4414
lemma continuous_closed_preimage:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4415
  assumes "continuous_on s f" "closed s" "closed t"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4416
  shows "closed {x \<in> s. f x \<in> t}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4417
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4418
  obtain T where T: "closed T" "{x \<in> s. f x \<in> t} = s \<inter> T"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4419
    using continuous_closed_in_preimage[OF assms(1,3)] unfolding closedin_closed by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4420
  thus ?thesis using closed_Int[of s T, OF assms(2)] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4421
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4422
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4423
lemma continuous_open_preimage_univ:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4424
  shows "\<forall>x. continuous (at x) f \<Longrightarrow> open s \<Longrightarrow> open {x. f x \<in> s}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4425
  using continuous_open_preimage[of UNIV f s] open_UNIV continuous_at_imp_continuous_on by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4426
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4427
lemma continuous_closed_preimage_univ:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4428
  shows "(\<forall>x. continuous (at x) f) \<Longrightarrow> closed s ==> closed {x. f x \<in> s}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4429
  using continuous_closed_preimage[of UNIV f s] closed_UNIV continuous_at_imp_continuous_on by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4430
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4431
lemma continuous_open_vimage:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4432
  shows "\<forall>x. continuous (at x) f \<Longrightarrow> open s \<Longrightarrow> open (f -` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4433
  unfolding vimage_def by (rule continuous_open_preimage_univ)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4434
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4435
lemma continuous_closed_vimage:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4436
  shows "\<forall>x. continuous (at x) f \<Longrightarrow> closed s \<Longrightarrow> closed (f -` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4437
  unfolding vimage_def by (rule continuous_closed_preimage_univ)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4438
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4439
lemma interior_image_subset:
35172
579dd5570f96 Added integration to Multivariate-Analysis (upto FTC)
himmelma
parents: 35028
diff changeset
  4440
  assumes "\<forall>x. continuous (at x) f" "inj f"
579dd5570f96 Added integration to Multivariate-Analysis (upto FTC)
himmelma
parents: 35028
diff changeset
  4441
  shows "interior (f ` s) \<subseteq> f ` (interior s)"
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  4442
proof
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  4443
  fix x assume "x \<in> interior (f ` s)"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  4444
  then obtain T where as: "open T" "x \<in> T" "T \<subseteq> f ` s" ..
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  4445
  hence "x \<in> f ` s" by auto
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  4446
  then obtain y where y: "y \<in> s" "x = f y" by auto
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  4447
  have "open (vimage f T)"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  4448
    using assms(1) `open T` by (rule continuous_open_vimage)
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  4449
  moreover have "y \<in> vimage f T"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  4450
    using `x = f y` `x \<in> T` by simp
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  4451
  moreover have "vimage f T \<subseteq> s"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  4452
    using `T \<subseteq> image f s` `inj f` unfolding inj_on_def subset_eq by auto
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  4453
  ultimately have "y \<in> interior s" ..
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  4454
  with `x = f y` show "x \<in> f ` interior s" ..
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  4455
qed
35172
579dd5570f96 Added integration to Multivariate-Analysis (upto FTC)
himmelma
parents: 35028
diff changeset
  4456
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  4457
text {* Equality of continuous functions on closure and related results. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4458
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4459
lemma continuous_closed_in_preimage_constant:
36668
941ba2da372e simplify definition of t1_space; generalize lemma closed_sing and related lemmas
huffman
parents: 36667
diff changeset
  4460
  fixes f :: "_ \<Rightarrow> 'b::t1_space"
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  4461
  shows "continuous_on s f ==> closedin (subtopology euclidean s) {x \<in> s. f x = a}"
36668
941ba2da372e simplify definition of t1_space; generalize lemma closed_sing and related lemmas
huffman
parents: 36667
diff changeset
  4462
  using continuous_closed_in_preimage[of s f "{a}"] by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4463
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4464
lemma continuous_closed_preimage_constant:
36668
941ba2da372e simplify definition of t1_space; generalize lemma closed_sing and related lemmas
huffman
parents: 36667
diff changeset
  4465
  fixes f :: "_ \<Rightarrow> 'b::t1_space"
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  4466
  shows "continuous_on s f \<Longrightarrow> closed s ==> closed {x \<in> s. f x = a}"
36668
941ba2da372e simplify definition of t1_space; generalize lemma closed_sing and related lemmas
huffman
parents: 36667
diff changeset
  4467
  using continuous_closed_preimage[of s f "{a}"] by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4468
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4469
lemma continuous_constant_on_closure:
36668
941ba2da372e simplify definition of t1_space; generalize lemma closed_sing and related lemmas
huffman
parents: 36667
diff changeset
  4470
  fixes f :: "_ \<Rightarrow> 'b::t1_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4471
  assumes "continuous_on (closure s) f"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4472
          "\<forall>x \<in> s. f x = a"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4473
  shows "\<forall>x \<in> (closure s). f x = a"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4474
    using continuous_closed_preimage_constant[of "closure s" f a]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4475
    assms closure_minimal[of s "{x \<in> closure s. f x = a}"] closure_subset unfolding subset_eq by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4476
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4477
lemma image_closure_subset:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4478
  assumes "continuous_on (closure s) f"  "closed t"  "(f ` s) \<subseteq> t"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4479
  shows "f ` (closure s) \<subseteq> t"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4480
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4481
  have "s \<subseteq> {x \<in> closure s. f x \<in> t}" using assms(3) closure_subset by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4482
  moreover have "closed {x \<in> closure s. f x \<in> t}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4483
    using continuous_closed_preimage[OF assms(1)] and assms(2) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4484
  ultimately have "closure s = {x \<in> closure s . f x \<in> t}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4485
    using closure_minimal[of s "{x \<in> closure s. f x \<in> t}"] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4486
  thus ?thesis by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4487
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4488
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4489
lemma continuous_on_closure_norm_le:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4490
  fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4491
  assumes "continuous_on (closure s) f"  "\<forall>y \<in> s. norm(f y) \<le> b"  "x \<in> (closure s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4492
  shows "norm(f x) \<le> b"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4493
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4494
  have *:"f ` s \<subseteq> cball 0 b" using assms(2)[unfolded mem_cball_0[THEN sym]] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4495
  show ?thesis
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4496
    using image_closure_subset[OF assms(1) closed_cball[of 0 b] *] assms(3)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4497
    unfolding subset_eq apply(erule_tac x="f x" in ballE) by (auto simp add: dist_norm)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4498
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4499
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  4500
text {* Making a continuous function avoid some value in a neighbourhood. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4501
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4502
lemma continuous_within_avoid:
50898
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4503
  fixes f :: "'a::metric_space \<Rightarrow> 'b::t1_space"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4504
  assumes "continuous (at x within s) f" and "f x \<noteq> a"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4505
  shows "\<exists>e>0. \<forall>y \<in> s. dist x y < e --> f y \<noteq> a"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4506
proof-
50898
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4507
  obtain U where "open U" and "f x \<in> U" and "a \<notin> U"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4508
    using t1_space [OF `f x \<noteq> a`] by fast
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4509
  have "(f ---> f x) (at x within s)"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4510
    using assms(1) by (simp add: continuous_within)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4511
  hence "eventually (\<lambda>y. f y \<in> U) (at x within s)"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4512
    using `open U` and `f x \<in> U`
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4513
    unfolding tendsto_def by fast
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4514
  hence "eventually (\<lambda>y. f y \<noteq> a) (at x within s)"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4515
    using `a \<notin> U` by (fast elim: eventually_mono [rotated])
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4516
  thus ?thesis
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4517
    unfolding Limits.eventually_within Limits.eventually_at
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4518
    by (rule ex_forward, cut_tac `f x \<noteq> a`, auto simp: dist_commute)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4519
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4520
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4521
lemma continuous_at_avoid:
50898
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4522
  fixes f :: "'a::metric_space \<Rightarrow> 'b::t1_space"
45031
9583f2b56f85 add lemmas within_empty and tendsto_bot;
huffman
parents: 44909
diff changeset
  4523
  assumes "continuous (at x) f" and "f x \<noteq> a"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4524
  shows "\<exists>e>0. \<forall>y. dist x y < e \<longrightarrow> f y \<noteq> a"
45031
9583f2b56f85 add lemmas within_empty and tendsto_bot;
huffman
parents: 44909
diff changeset
  4525
  using assms continuous_within_avoid[of x UNIV f a] by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4526
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4527
lemma continuous_on_avoid:
50898
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4528
  fixes f :: "'a::metric_space \<Rightarrow> 'b::t1_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4529
  assumes "continuous_on s f"  "x \<in> s"  "f x \<noteq> a"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4530
  shows "\<exists>e>0. \<forall>y \<in> s. dist x y < e \<longrightarrow> f y \<noteq> a"
50898
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4531
using assms(1)[unfolded continuous_on_eq_continuous_within, THEN bspec[where x=x], OF assms(2)]  continuous_within_avoid[of x s f a]  assms(3) by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4532
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4533
lemma continuous_on_open_avoid:
50898
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4534
  fixes f :: "'a::metric_space \<Rightarrow> 'b::t1_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4535
  assumes "continuous_on s f"  "open s"  "x \<in> s"  "f x \<noteq> a"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4536
  shows "\<exists>e>0. \<forall>y. dist x y < e \<longrightarrow> f y \<noteq> a"
50898
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4537
using assms(1)[unfolded continuous_on_eq_continuous_at[OF assms(2)], THEN bspec[where x=x], OF assms(3)]  continuous_at_avoid[of x f a]  assms(4) by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4538
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  4539
text {* Proving a function is constant by proving open-ness of level set. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4540
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4541
lemma continuous_levelset_open_in_cases:
36668
941ba2da372e simplify definition of t1_space; generalize lemma closed_sing and related lemmas
huffman
parents: 36667
diff changeset
  4542
  fixes f :: "_ \<Rightarrow> 'b::t1_space"
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  4543
  shows "connected s \<Longrightarrow> continuous_on s f \<Longrightarrow>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4544
        openin (subtopology euclidean s) {x \<in> s. f x = a}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4545
        ==> (\<forall>x \<in> s. f x \<noteq> a) \<or> (\<forall>x \<in> s. f x = a)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4546
unfolding connected_clopen using continuous_closed_in_preimage_constant by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4547
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4548
lemma continuous_levelset_open_in:
36668
941ba2da372e simplify definition of t1_space; generalize lemma closed_sing and related lemmas
huffman
parents: 36667
diff changeset
  4549
  fixes f :: "_ \<Rightarrow> 'b::t1_space"
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  4550
  shows "connected s \<Longrightarrow> continuous_on s f \<Longrightarrow>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4551
        openin (subtopology euclidean s) {x \<in> s. f x = a} \<Longrightarrow>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4552
        (\<exists>x \<in> s. f x = a)  ==> (\<forall>x \<in> s. f x = a)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4553
using continuous_levelset_open_in_cases[of s f ]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4554
by meson
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4555
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4556
lemma continuous_levelset_open:
36668
941ba2da372e simplify definition of t1_space; generalize lemma closed_sing and related lemmas
huffman
parents: 36667
diff changeset
  4557
  fixes f :: "_ \<Rightarrow> 'b::t1_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4558
  assumes "connected s"  "continuous_on s f"  "open {x \<in> s. f x = a}"  "\<exists>x \<in> s.  f x = a"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4559
  shows "\<forall>x \<in> s. f x = a"
36362
06475a1547cb fix lots of looping simp calls and other warnings
huffman
parents: 36360
diff changeset
  4560
using continuous_levelset_open_in[OF assms(1,2), of a, unfolded openin_open] using assms (3,4) by fast
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4561
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  4562
text {* Some arithmetical combinations (more to prove). *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4563
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4564
lemma open_scaling[intro]:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4565
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4566
  assumes "c \<noteq> 0"  "open s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4567
  shows "open((\<lambda>x. c *\<^sub>R x) ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4568
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4569
  { fix x assume "x \<in> s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4570
    then obtain e where "e>0" and e:"\<forall>x'. dist x' x < e \<longrightarrow> x' \<in> s" using assms(2)[unfolded open_dist, THEN bspec[where x=x]] by auto
36778
739a9379e29b avoid using real-specific versions of generic lemmas
huffman
parents: 36669
diff changeset
  4571
    have "e * abs c > 0" using assms(1)[unfolded zero_less_abs_iff[THEN sym]] using mult_pos_pos[OF `e>0`] by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4572
    moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4573
    { fix y assume "dist y (c *\<^sub>R x) < e * \<bar>c\<bar>"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4574
      hence "norm ((1 / c) *\<^sub>R y - x) < e" unfolding dist_norm
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4575
        using norm_scaleR[of c "(1 / c) *\<^sub>R y - x", unfolded scaleR_right_diff_distrib, unfolded scaleR_scaleR] assms(1)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4576
          assms(1)[unfolded zero_less_abs_iff[THEN sym]] by (simp del:zero_less_abs_iff)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4577
      hence "y \<in> op *\<^sub>R c ` s" using rev_image_eqI[of "(1 / c) *\<^sub>R y" s y "op *\<^sub>R c"]  e[THEN spec[where x="(1 / c) *\<^sub>R y"]]  assms(1) unfolding dist_norm scaleR_scaleR by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4578
    ultimately have "\<exists>e>0. \<forall>x'. dist x' (c *\<^sub>R x) < e \<longrightarrow> x' \<in> op *\<^sub>R c ` s" apply(rule_tac x="e * abs c" in exI) by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4579
  thus ?thesis unfolding open_dist by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4580
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4581
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4582
lemma minus_image_eq_vimage:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4583
  fixes A :: "'a::ab_group_add set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4584
  shows "(\<lambda>x. - x) ` A = (\<lambda>x. - x) -` A"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4585
  by (auto intro!: image_eqI [where f="\<lambda>x. - x"])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4586
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4587
lemma open_negations:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4588
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4589
  shows "open s ==> open ((\<lambda> x. -x) ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4590
  unfolding scaleR_minus1_left [symmetric]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4591
  by (rule open_scaling, auto)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4592
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4593
lemma open_translation:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4594
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4595
  assumes "open s"  shows "open((\<lambda>x. a + x) ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4596
proof-
44647
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4597
  { fix x have "continuous (at x) (\<lambda>x. x - a)"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4598
      by (intro continuous_diff continuous_at_id continuous_const) }
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4599
  moreover have "{x. x - a \<in> s} = op + a ` s" by force
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4600
  ultimately show ?thesis using continuous_open_preimage_univ[of "\<lambda>x. x - a" s] using assms by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4601
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4602
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4603
lemma open_affinity:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4604
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4605
  assumes "open s"  "c \<noteq> 0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4606
  shows "open ((\<lambda>x. a + c *\<^sub>R x) ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4607
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4608
  have *:"(\<lambda>x. a + c *\<^sub>R x) = (\<lambda>x. a + x) \<circ> (\<lambda>x. c *\<^sub>R x)" unfolding o_def ..
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4609
  have "op + a ` op *\<^sub>R c ` s = (op + a \<circ> op *\<^sub>R c) ` s" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4610
  thus ?thesis using assms open_translation[of "op *\<^sub>R c ` s" a] unfolding * by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4611
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4612
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4613
lemma interior_translation:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4614
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4615
  shows "interior ((\<lambda>x. a + x) ` s) = (\<lambda>x. a + x) ` (interior s)"
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
  4616
proof (rule set_eqI, rule)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4617
  fix x assume "x \<in> interior (op + a ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4618
  then obtain e where "e>0" and e:"ball x e \<subseteq> op + a ` s" unfolding mem_interior by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4619
  hence "ball (x - a) e \<subseteq> s" unfolding subset_eq Ball_def mem_ball dist_norm apply auto apply(erule_tac x="a + xa" in allE) unfolding ab_group_add_class.diff_diff_eq[THEN sym] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4620
  thus "x \<in> op + a ` interior s" unfolding image_iff apply(rule_tac x="x - a" in bexI) unfolding mem_interior using `e > 0` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4621
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4622
  fix x assume "x \<in> op + a ` interior s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4623
  then obtain y e where "e>0" and e:"ball y e \<subseteq> s" and y:"x = a + y" unfolding image_iff Bex_def mem_interior by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4624
  { fix z have *:"a + y - z = y + a - z" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4625
    assume "z\<in>ball x e"
45548
3e2722d66169 Groups.thy: generalize several lemmas from class ab_group_add to class group_add
huffman
parents: 45270
diff changeset
  4626
    hence "z - a \<in> s" using e[unfolded subset_eq, THEN bspec[where x="z - a"]] unfolding mem_ball dist_norm y group_add_class.diff_diff_eq2 * by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4627
    hence "z \<in> op + a ` s" unfolding image_iff by(auto intro!: bexI[where x="z - a"])  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4628
  hence "ball x e \<subseteq> op + a ` s" unfolding subset_eq by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4629
  thus "x \<in> interior (op + a ` s)" unfolding mem_interior using `e>0` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4630
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4631
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4632
text {* Topological properties of linear functions. *}
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4633
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4634
lemma linear_lim_0:
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4635
  assumes "bounded_linear f" shows "(f ---> 0) (at (0))"
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4636
proof-
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4637
  interpret f: bounded_linear f by fact
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4638
  have "(f ---> f 0) (at 0)"
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4639
    using tendsto_ident_at by (rule f.tendsto)
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4640
  thus ?thesis unfolding f.zero .
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4641
qed
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4642
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4643
lemma linear_continuous_at:
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4644
  assumes "bounded_linear f"  shows "continuous (at a) f"
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4645
  unfolding continuous_at using assms
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4646
  apply (rule bounded_linear.tendsto)
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4647
  apply (rule tendsto_ident_at)
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4648
  done
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4649
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4650
lemma linear_continuous_within:
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4651
  shows "bounded_linear f ==> continuous (at x within s) f"
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4652
  using continuous_at_imp_continuous_within[of x f s] using linear_continuous_at[of f] by auto
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4653
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4654
lemma linear_continuous_on:
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4655
  shows "bounded_linear f ==> continuous_on s f"
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4656
  using continuous_at_imp_continuous_on[of s f] using linear_continuous_at[of f] by auto
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4657
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  4658
text {* Also bilinear functions, in composition form. *}
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4659
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4660
lemma bilinear_continuous_at_compose:
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4661
  shows "continuous (at x) f \<Longrightarrow> continuous (at x) g \<Longrightarrow> bounded_bilinear h
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4662
        ==> continuous (at x) (\<lambda>x. h (f x) (g x))"
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4663
  unfolding continuous_at using Lim_bilinear[of f "f x" "(at x)" g "g x" h] by auto
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4664
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4665
lemma bilinear_continuous_within_compose:
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4666
  shows "continuous (at x within s) f \<Longrightarrow> continuous (at x within s) g \<Longrightarrow> bounded_bilinear h
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4667
        ==> continuous (at x within s) (\<lambda>x. h (f x) (g x))"
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4668
  unfolding continuous_within using Lim_bilinear[of f "f x"] by auto
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4669
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4670
lemma bilinear_continuous_on_compose:
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4671
  shows "continuous_on s f \<Longrightarrow> continuous_on s g \<Longrightarrow> bounded_bilinear h
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4672
             ==> continuous_on s (\<lambda>x. h (f x) (g x))"
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4673
  unfolding continuous_on_def
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4674
  by (fast elim: bounded_bilinear.tendsto)
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4675
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  4676
text {* Preservation of compactness and connectedness under continuous function. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4677
50898
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4678
lemma compact_eq_openin_cover:
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4679
  "compact S \<longleftrightarrow>
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4680
    (\<forall>C. (\<forall>c\<in>C. openin (subtopology euclidean S) c) \<and> S \<subseteq> \<Union>C \<longrightarrow>
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4681
      (\<exists>D\<subseteq>C. finite D \<and> S \<subseteq> \<Union>D))"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4682
proof safe
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4683
  fix C
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4684
  assume "compact S" and "\<forall>c\<in>C. openin (subtopology euclidean S) c" and "S \<subseteq> \<Union>C"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4685
  hence "\<forall>c\<in>{T. open T \<and> S \<inter> T \<in> C}. open c" and "S \<subseteq> \<Union>{T. open T \<and> S \<inter> T \<in> C}"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4686
    unfolding openin_open by force+
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4687
  with `compact S` obtain D where "D \<subseteq> {T. open T \<and> S \<inter> T \<in> C}" and "finite D" and "S \<subseteq> \<Union>D"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4688
    by (rule compactE)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4689
  hence "image (\<lambda>T. S \<inter> T) D \<subseteq> C \<and> finite (image (\<lambda>T. S \<inter> T) D) \<and> S \<subseteq> \<Union>(image (\<lambda>T. S \<inter> T) D)"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4690
    by auto
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4691
  thus "\<exists>D\<subseteq>C. finite D \<and> S \<subseteq> \<Union>D" ..
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4692
next
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4693
  assume 1: "\<forall>C. (\<forall>c\<in>C. openin (subtopology euclidean S) c) \<and> S \<subseteq> \<Union>C \<longrightarrow>
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4694
        (\<exists>D\<subseteq>C. finite D \<and> S \<subseteq> \<Union>D)"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4695
  show "compact S"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4696
  proof (rule compactI)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4697
    fix C
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4698
    let ?C = "image (\<lambda>T. S \<inter> T) C"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4699
    assume "\<forall>t\<in>C. open t" and "S \<subseteq> \<Union>C"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4700
    hence "(\<forall>c\<in>?C. openin (subtopology euclidean S) c) \<and> S \<subseteq> \<Union>?C"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4701
      unfolding openin_open by auto
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4702
    with 1 obtain D where "D \<subseteq> ?C" and "finite D" and "S \<subseteq> \<Union>D"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4703
      by metis
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4704
    let ?D = "inv_into C (\<lambda>T. S \<inter> T) ` D"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4705
    have "?D \<subseteq> C \<and> finite ?D \<and> S \<subseteq> \<Union>?D"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4706
    proof (intro conjI)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4707
      from `D \<subseteq> ?C` show "?D \<subseteq> C"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4708
        by (fast intro: inv_into_into)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4709
      from `finite D` show "finite ?D"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4710
        by (rule finite_imageI)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4711
      from `S \<subseteq> \<Union>D` show "S \<subseteq> \<Union>?D"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4712
        apply (rule subset_trans)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4713
        apply clarsimp
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4714
        apply (frule subsetD [OF `D \<subseteq> ?C`, THEN f_inv_into_f])
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4715
        apply (erule rev_bexI, fast)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4716
        done
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4717
    qed
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4718
    thus "\<exists>D\<subseteq>C. finite D \<and> S \<subseteq> \<Union>D" ..
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4719
  qed
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4720
qed
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4721
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4722
lemma compact_continuous_image:
50898
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4723
  assumes "continuous_on s f" and "compact s"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4724
  shows "compact (f ` s)"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4725
using assms (* FIXME: long unstructured proof *)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4726
unfolding continuous_on_open
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4727
unfolding compact_eq_openin_cover
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4728
apply clarify
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4729
apply (drule_tac x="image (\<lambda>t. {x \<in> s. f x \<in> t}) C" in spec)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4730
apply (drule mp)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4731
apply (rule conjI)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4732
apply simp
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4733
apply clarsimp
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4734
apply (drule subsetD)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4735
apply (erule imageI)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4736
apply fast
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4737
apply (erule thin_rl)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4738
apply clarify
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4739
apply (rule_tac x="image (inv_into C (\<lambda>t. {x \<in> s. f x \<in> t})) D" in exI)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4740
apply (intro conjI)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4741
apply clarify
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4742
apply (rule inv_into_into)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4743
apply (erule (1) subsetD)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4744
apply (erule finite_imageI)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4745
apply (clarsimp, rename_tac x)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4746
apply (drule (1) subsetD, clarify)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4747
apply (drule (1) subsetD, clarify)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4748
apply (rule rev_bexI)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4749
apply assumption
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4750
apply (subgoal_tac "{x \<in> s. f x \<in> t} \<in> (\<lambda>t. {x \<in> s. f x \<in> t}) ` C")
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4751
apply (drule f_inv_into_f)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4752
apply fast
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4753
apply (erule imageI)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4754
done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4755
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4756
lemma connected_continuous_image:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4757
  assumes "continuous_on s f"  "connected s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4758
  shows "connected(f ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4759
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4760
  { fix T assume as: "T \<noteq> {}"  "T \<noteq> f ` s"  "openin (subtopology euclidean (f ` s)) T"  "closedin (subtopology euclidean (f ` s)) T"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4761
    have "{x \<in> s. f x \<in> T} = {} \<or> {x \<in> s. f x \<in> T} = s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4762
      using assms(1)[unfolded continuous_on_open, THEN spec[where x=T]]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4763
      using assms(1)[unfolded continuous_on_closed, THEN spec[where x=T]]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4764
      using assms(2)[unfolded connected_clopen, THEN spec[where x="{x \<in> s. f x \<in> T}"]] as(3,4) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4765
    hence False using as(1,2)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4766
      using as(4)[unfolded closedin_def topspace_euclidean_subtopology] by auto }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4767
  thus ?thesis unfolding connected_clopen by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4768
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4769
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  4770
text {* Continuity implies uniform continuity on a compact domain. *}
50943
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  4771
  
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4772
lemma compact_uniformly_continuous:
50943
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  4773
  assumes f: "continuous_on s f" and s: "compact s"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4774
  shows "uniformly_continuous_on s f"
50943
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  4775
  unfolding uniformly_continuous_on_def
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  4776
proof (cases, safe)
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  4777
  fix e :: real assume "0 < e" "s \<noteq> {}"
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  4778
  def [simp]: R \<equiv> "{(y, d). y \<in> s \<and> 0 < d \<and> ball y d \<inter> s \<subseteq> {x \<in> s. f x \<in> ball (f y) (e/2) } }"
50944
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  4779
  let ?b = "(\<lambda>(y, d). ball y (d/2))"
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  4780
  have "(\<forall>r\<in>R. open (?b r))" "s \<subseteq> (\<Union>r\<in>R. ?b r)"
50943
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  4781
  proof safe
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  4782
    fix y assume "y \<in> s"
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  4783
    from continuous_open_in_preimage[OF f open_ball]
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  4784
    obtain T where "open T" and T: "{x \<in> s. f x \<in> ball (f y) (e/2)} = T \<inter> s"
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  4785
      unfolding openin_subtopology open_openin by metis
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  4786
    then obtain d where "ball y d \<subseteq> T" "0 < d"
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  4787
      using `0 < e` `y \<in> s` by (auto elim!: openE)
50944
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  4788
    with T `y \<in> s` show "y \<in> (\<Union>r\<in>R. ?b r)"
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  4789
      by (intro UN_I[of "(y, d)"]) auto
50943
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  4790
  qed auto
50944
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  4791
  with s obtain D where D: "finite D" "D \<subseteq> R" "s \<subseteq> (\<Union>(y, d)\<in>D. ball y (d/2))"
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  4792
    by (rule compactE_image)
50943
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  4793
  with `s \<noteq> {}` have [simp]: "\<And>x. x < Min (snd ` D) \<longleftrightarrow> (\<forall>(y, d)\<in>D. x < d)"
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  4794
    by (subst Min_gr_iff) auto
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  4795
  show "\<exists>d>0. \<forall>x\<in>s. \<forall>x'\<in>s. dist x' x < d \<longrightarrow> dist (f x') (f x) < e"
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  4796
  proof (rule, safe)
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  4797
    fix x x' assume in_s: "x' \<in> s" "x \<in> s"
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  4798
    with D obtain y d where x: "x \<in> ball y (d/2)" "(y, d) \<in> D"
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  4799
      by blast
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  4800
    moreover assume "dist x x' < Min (snd`D) / 2"
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  4801
    ultimately have "dist y x' < d"
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  4802
      by (intro dist_double[where x=x and d=d]) (auto simp: dist_commute)
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  4803
    with D x in_s show  "dist (f x) (f x') < e"
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  4804
      by (intro dist_double[where x="f y" and d=e]) (auto simp: dist_commute subset_eq)
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  4805
  qed (insert D, auto)
88a00a1c7c2c use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents: 50942
diff changeset
  4806
qed auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4807
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4808
text{* Continuity of inverse function on compact domain. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4809
44647
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4810
lemma continuous_on_inv:
50898
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4811
  fixes f :: "'a::topological_space \<Rightarrow> 'b::t2_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4812
  assumes "continuous_on s f"  "compact s"  "\<forall>x \<in> s. g (f x) = x"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4813
  shows "continuous_on (f ` s) g"
50898
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4814
unfolding continuous_on_topological
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4815
proof (clarsimp simp add: assms(3))
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4816
  fix x :: 'a and B :: "'a set"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4817
  assume "x \<in> s" and "open B" and "x \<in> B"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4818
  have 1: "\<forall>x\<in>s. f x \<in> f ` (s - B) \<longleftrightarrow> x \<in> s - B"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4819
    using assms(3) by (auto, metis)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4820
  have "continuous_on (s - B) f"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4821
    using `continuous_on s f` Diff_subset
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4822
    by (rule continuous_on_subset)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4823
  moreover have "compact (s - B)"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4824
    using `open B` and `compact s`
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4825
    unfolding Diff_eq by (intro compact_inter_closed closed_Compl)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4826
  ultimately have "compact (f ` (s - B))"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4827
    by (rule compact_continuous_image)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4828
  hence "closed (f ` (s - B))"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4829
    by (rule compact_imp_closed)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4830
  hence "open (- f ` (s - B))"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4831
    by (rule open_Compl)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4832
  moreover have "f x \<in> - f ` (s - B)"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4833
    using `x \<in> s` and `x \<in> B` by (simp add: 1)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4834
  moreover have "\<forall>y\<in>s. f y \<in> - f ` (s - B) \<longrightarrow> y \<in> B"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4835
    by (simp add: 1)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4836
  ultimately show "\<exists>A. open A \<and> f x \<in> A \<and> (\<forall>y\<in>s. f y \<in> A \<longrightarrow> y \<in> B)"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4837
    by fast
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4838
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4839
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4840
text {* A uniformly convergent limit of continuous functions is continuous. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4841
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4842
lemma continuous_uniform_limit:
44212
4d10e7f342b1 generalize lemma continuous_uniform_limit to class metric_space
huffman
parents: 44211
diff changeset
  4843
  fixes f :: "'a \<Rightarrow> 'b::metric_space \<Rightarrow> 'c::metric_space"
4d10e7f342b1 generalize lemma continuous_uniform_limit to class metric_space
huffman
parents: 44211
diff changeset
  4844
  assumes "\<not> trivial_limit F"
4d10e7f342b1 generalize lemma continuous_uniform_limit to class metric_space
huffman
parents: 44211
diff changeset
  4845
  assumes "eventually (\<lambda>n. continuous_on s (f n)) F"
4d10e7f342b1 generalize lemma continuous_uniform_limit to class metric_space
huffman
parents: 44211
diff changeset
  4846
  assumes "\<forall>e>0. eventually (\<lambda>n. \<forall>x\<in>s. dist (f n x) (g x) < e) F"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4847
  shows "continuous_on s g"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4848
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4849
  { fix x and e::real assume "x\<in>s" "e>0"
44212
4d10e7f342b1 generalize lemma continuous_uniform_limit to class metric_space
huffman
parents: 44211
diff changeset
  4850
    have "eventually (\<lambda>n. \<forall>x\<in>s. dist (f n x) (g x) < e / 3) F"
4d10e7f342b1 generalize lemma continuous_uniform_limit to class metric_space
huffman
parents: 44211
diff changeset
  4851
      using `e>0` assms(3)[THEN spec[where x="e/3"]] by auto
4d10e7f342b1 generalize lemma continuous_uniform_limit to class metric_space
huffman
parents: 44211
diff changeset
  4852
    from eventually_happens [OF eventually_conj [OF this assms(2)]]
4d10e7f342b1 generalize lemma continuous_uniform_limit to class metric_space
huffman
parents: 44211
diff changeset
  4853
    obtain n where n:"\<forall>x\<in>s. dist (f n x) (g x) < e / 3"  "continuous_on s (f n)"
4d10e7f342b1 generalize lemma continuous_uniform_limit to class metric_space
huffman
parents: 44211
diff changeset
  4854
      using assms(1) by blast
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4855
    have "e / 3 > 0" using `e>0` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4856
    then obtain d where "d>0" and d:"\<forall>x'\<in>s. dist x' x < d \<longrightarrow> dist (f n x') (f n x) < e / 3"
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  4857
      using n(2)[unfolded continuous_on_iff, THEN bspec[where x=x], OF `x\<in>s`, THEN spec[where x="e/3"]] by blast
44212
4d10e7f342b1 generalize lemma continuous_uniform_limit to class metric_space
huffman
parents: 44211
diff changeset
  4858
    { fix y assume "y \<in> s" and "dist y x < d"
4d10e7f342b1 generalize lemma continuous_uniform_limit to class metric_space
huffman
parents: 44211
diff changeset
  4859
      hence "dist (f n y) (f n x) < e / 3"
4d10e7f342b1 generalize lemma continuous_uniform_limit to class metric_space
huffman
parents: 44211
diff changeset
  4860
        by (rule d [rule_format])
4d10e7f342b1 generalize lemma continuous_uniform_limit to class metric_space
huffman
parents: 44211
diff changeset
  4861
      hence "dist (f n y) (g x) < 2 * e / 3"
4d10e7f342b1 generalize lemma continuous_uniform_limit to class metric_space
huffman
parents: 44211
diff changeset
  4862
        using dist_triangle [of "f n y" "g x" "f n x"]
4d10e7f342b1 generalize lemma continuous_uniform_limit to class metric_space
huffman
parents: 44211
diff changeset
  4863
        using n(1)[THEN bspec[where x=x], OF `x\<in>s`]
4d10e7f342b1 generalize lemma continuous_uniform_limit to class metric_space
huffman
parents: 44211
diff changeset
  4864
        by auto
4d10e7f342b1 generalize lemma continuous_uniform_limit to class metric_space
huffman
parents: 44211
diff changeset
  4865
      hence "dist (g y) (g x) < e"
4d10e7f342b1 generalize lemma continuous_uniform_limit to class metric_space
huffman
parents: 44211
diff changeset
  4866
        using n(1)[THEN bspec[where x=y], OF `y\<in>s`]
4d10e7f342b1 generalize lemma continuous_uniform_limit to class metric_space
huffman
parents: 44211
diff changeset
  4867
        using dist_triangle3 [of "g y" "g x" "f n y"]
4d10e7f342b1 generalize lemma continuous_uniform_limit to class metric_space
huffman
parents: 44211
diff changeset
  4868
        by auto }
4d10e7f342b1 generalize lemma continuous_uniform_limit to class metric_space
huffman
parents: 44211
diff changeset
  4869
    hence "\<exists>d>0. \<forall>x'\<in>s. dist x' x < d \<longrightarrow> dist (g x') (g x) < e"
4d10e7f342b1 generalize lemma continuous_uniform_limit to class metric_space
huffman
parents: 44211
diff changeset
  4870
      using `d>0` by auto }
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  4871
  thus ?thesis unfolding continuous_on_iff by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4872
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4873
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  4874
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  4875
subsection {* Topological stuff lifted from and dropped to R *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4876
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4877
lemma open_real:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4878
  fixes s :: "real set" shows
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4879
 "open s \<longleftrightarrow>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4880
        (\<forall>x \<in> s. \<exists>e>0. \<forall>x'. abs(x' - x) < e --> x' \<in> s)" (is "?lhs = ?rhs")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4881
  unfolding open_dist dist_norm by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4882
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4883
lemma islimpt_approachable_real:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4884
  fixes s :: "real set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4885
  shows "x islimpt s \<longleftrightarrow> (\<forall>e>0.  \<exists>x'\<in> s. x' \<noteq> x \<and> abs(x' - x) < e)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4886
  unfolding islimpt_approachable dist_norm by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4887
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4888
lemma closed_real:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4889
  fixes s :: "real set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4890
  shows "closed s \<longleftrightarrow>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4891
        (\<forall>x. (\<forall>e>0.  \<exists>x' \<in> s. x' \<noteq> x \<and> abs(x' - x) < e)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4892
            --> x \<in> s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4893
  unfolding closed_limpt islimpt_approachable dist_norm by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4894
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4895
lemma continuous_at_real_range:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4896
  fixes f :: "'a::real_normed_vector \<Rightarrow> real"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4897
  shows "continuous (at x) f \<longleftrightarrow> (\<forall>e>0. \<exists>d>0.
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4898
        \<forall>x'. norm(x' - x) < d --> abs(f x' - f x) < e)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4899
  unfolding continuous_at unfolding Lim_at
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4900
  unfolding dist_nz[THEN sym] unfolding dist_norm apply auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4901
  apply(erule_tac x=e in allE) apply auto apply (rule_tac x=d in exI) apply auto apply (erule_tac x=x' in allE) apply auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4902
  apply(erule_tac x=e in allE) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4903
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4904
lemma continuous_on_real_range:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4905
  fixes f :: "'a::real_normed_vector \<Rightarrow> real"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4906
  shows "continuous_on s f \<longleftrightarrow> (\<forall>x \<in> s. \<forall>e>0. \<exists>d>0. (\<forall>x' \<in> s. norm(x' - x) < d --> abs(f x' - f x) < e))"
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  4907
  unfolding continuous_on_iff dist_norm by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4908
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  4909
text {* Hence some handy theorems on distance, diameter etc. of/from a set. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4910
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4911
lemma compact_attains_sup:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4912
  fixes s :: "real set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4913
  assumes "compact s"  "s \<noteq> {}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4914
  shows "\<exists>x \<in> s. \<forall>y \<in> s. y \<le> x"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4915
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4916
  from assms(1) have a:"bounded s" "closed s" unfolding compact_eq_bounded_closed by auto
33270
paulson
parents: 33175
diff changeset
  4917
  { fix e::real assume as: "\<forall>x\<in>s. x \<le> Sup s" "Sup s \<notin> s"  "0 < e" "\<forall>x'\<in>s. x' = Sup s \<or> \<not> Sup s - x' < e"
paulson
parents: 33175
diff changeset
  4918
    have "isLub UNIV s (Sup s)" using Sup[OF assms(2)] unfolding setle_def using as(1) by auto
paulson
parents: 33175
diff changeset
  4919
    moreover have "isUb UNIV s (Sup s - e)" unfolding isUb_def unfolding setle_def using as(4,2) by auto
paulson
parents: 33175
diff changeset
  4920
    ultimately have False using isLub_le_isUb[of UNIV s "Sup s" "Sup s - e"] using `e>0` by auto  }
paulson
parents: 33175
diff changeset
  4921
  thus ?thesis using bounded_has_Sup(1)[OF a(1) assms(2)] using a(2)[unfolded closed_real, THEN spec[where x="Sup s"]]
paulson
parents: 33175
diff changeset
  4922
    apply(rule_tac x="Sup s" in bexI) by auto
paulson
parents: 33175
diff changeset
  4923
qed
paulson
parents: 33175
diff changeset
  4924
paulson
parents: 33175
diff changeset
  4925
lemma Inf:
paulson
parents: 33175
diff changeset
  4926
  fixes S :: "real set"
paulson
parents: 33175
diff changeset
  4927
  shows "S \<noteq> {} ==> (\<exists>b. b <=* S) ==> isGlb UNIV S (Inf S)"
paulson
parents: 33175
diff changeset
  4928
by (auto simp add: isLb_def setle_def setge_def isGlb_def greatestP_def) 
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4929
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4930
lemma compact_attains_inf:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4931
  fixes s :: "real set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4932
  assumes "compact s" "s \<noteq> {}"  shows "\<exists>x \<in> s. \<forall>y \<in> s. x \<le> y"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4933
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4934
  from assms(1) have a:"bounded s" "closed s" unfolding compact_eq_bounded_closed by auto
33270
paulson
parents: 33175
diff changeset
  4935
  { fix e::real assume as: "\<forall>x\<in>s. x \<ge> Inf s"  "Inf s \<notin> s"  "0 < e"
paulson
parents: 33175
diff changeset
  4936
      "\<forall>x'\<in>s. x' = Inf s \<or> \<not> abs (x' - Inf s) < e"
paulson
parents: 33175
diff changeset
  4937
    have "isGlb UNIV s (Inf s)" using Inf[OF assms(2)] unfolding setge_def using as(1) by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4938
    moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4939
    { fix x assume "x \<in> s"
33270
paulson
parents: 33175
diff changeset
  4940
      hence *:"abs (x - Inf s) = x - Inf s" using as(1)[THEN bspec[where x=x]] by auto
paulson
parents: 33175
diff changeset
  4941
      have "Inf s + e \<le> x" using as(4)[THEN bspec[where x=x]] using as(2) `x\<in>s` unfolding * by auto }
paulson
parents: 33175
diff changeset
  4942
    hence "isLb UNIV s (Inf s + e)" unfolding isLb_def and setge_def by auto
paulson
parents: 33175
diff changeset
  4943
    ultimately have False using isGlb_le_isLb[of UNIV s "Inf s" "Inf s + e"] using `e>0` by auto  }
paulson
parents: 33175
diff changeset
  4944
  thus ?thesis using bounded_has_Inf(1)[OF a(1) assms(2)] using a(2)[unfolded closed_real, THEN spec[where x="Inf s"]]
paulson
parents: 33175
diff changeset
  4945
    apply(rule_tac x="Inf s" in bexI) by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4946
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4947
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4948
lemma continuous_attains_sup:
50948
8c742f9de9f5 generalize topology lemmas; simplify proofs
huffman
parents: 50944
diff changeset
  4949
  fixes f :: "'a::topological_space \<Rightarrow> real"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4950
  shows "compact s \<Longrightarrow> s \<noteq> {} \<Longrightarrow> continuous_on s f
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4951
        ==> (\<exists>x \<in> s. \<forall>y \<in> s.  f y \<le> f x)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4952
  using compact_attains_sup[of "f ` s"]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4953
  using compact_continuous_image[of s f] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4954
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4955
lemma continuous_attains_inf:
50948
8c742f9de9f5 generalize topology lemmas; simplify proofs
huffman
parents: 50944
diff changeset
  4956
  fixes f :: "'a::topological_space \<Rightarrow> real"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4957
  shows "compact s \<Longrightarrow> s \<noteq> {} \<Longrightarrow> continuous_on s f
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4958
        \<Longrightarrow> (\<exists>x \<in> s. \<forall>y \<in> s. f x \<le> f y)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4959
  using compact_attains_inf[of "f ` s"]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4960
  using compact_continuous_image[of s f] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4961
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4962
lemma distance_attains_sup:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4963
  assumes "compact s" "s \<noteq> {}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4964
  shows "\<exists>x \<in> s. \<forall>y \<in> s. dist a y \<le> dist a x"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4965
proof (rule continuous_attains_sup [OF assms])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4966
  { fix x assume "x\<in>s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4967
    have "(dist a ---> dist a x) (at x within s)"
44568
e6f291cb5810 discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
huffman
parents: 44533
diff changeset
  4968
      by (intro tendsto_dist tendsto_const Lim_at_within tendsto_ident_at)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4969
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4970
  thus "continuous_on s (dist a)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4971
    unfolding continuous_on ..
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4972
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4973
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  4974
text {* For \emph{minimal} distance, we only need closure, not compactness. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4975
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4976
lemma distance_attains_inf:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4977
  fixes a :: "'a::heine_borel"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4978
  assumes "closed s"  "s \<noteq> {}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4979
  shows "\<exists>x \<in> s. \<forall>y \<in> s. dist a x \<le> dist a y"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4980
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4981
  from assms(2) obtain b where "b\<in>s" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4982
  let ?B = "cball a (dist b a) \<inter> s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4983
  have "b \<in> ?B" using `b\<in>s` by (simp add: dist_commute)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4984
  hence "?B \<noteq> {}" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4985
  moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4986
  { fix x assume "x\<in>?B"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4987
    fix e::real assume "e>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4988
    { fix x' assume "x'\<in>?B" and as:"dist x' x < e"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4989
      from as have "\<bar>dist a x' - dist a x\<bar> < e"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4990
        unfolding abs_less_iff minus_diff_eq
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4991
        using dist_triangle2 [of a x' x]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4992
        using dist_triangle [of a x x']
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4993
        by arith
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4994
    }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4995
    hence "\<exists>d>0. \<forall>x'\<in>?B. dist x' x < d \<longrightarrow> \<bar>dist a x' - dist a x\<bar> < e"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4996
      using `e>0` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4997
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4998
  hence "continuous_on (cball a (dist b a) \<inter> s) (dist a)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4999
    unfolding continuous_on Lim_within dist_norm real_norm_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5000
    by fast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5001
  moreover have "compact ?B"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5002
    using compact_cball[of a "dist b a"]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5003
    unfolding compact_eq_bounded_closed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5004
    using bounded_Int and closed_Int and assms(1) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5005
  ultimately obtain x where "x\<in>cball a (dist b a) \<inter> s" "\<forall>y\<in>cball a (dist b a) \<inter> s. dist a x \<le> dist a y"
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 44668
diff changeset
  5006
    using continuous_attains_inf[of ?B "dist a"] by fastforce
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 44668
diff changeset
  5007
  thus ?thesis by fastforce
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5008
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5009
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  5010
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  5011
subsection {* Pasted sets *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5012
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5013
lemma bounded_Times:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5014
  assumes "bounded s" "bounded t" shows "bounded (s \<times> t)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5015
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5016
  obtain x y a b where "\<forall>z\<in>s. dist x z \<le> a" "\<forall>z\<in>t. dist y z \<le> b"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5017
    using assms [unfolded bounded_def] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5018
  then have "\<forall>z\<in>s \<times> t. dist (x, y) z \<le> sqrt (a\<twosuperior> + b\<twosuperior>)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5019
    by (auto simp add: dist_Pair_Pair real_sqrt_le_mono add_mono power_mono)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5020
  thus ?thesis unfolding bounded_any_center [where a="(x, y)"] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5021
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5022
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5023
lemma mem_Times_iff: "x \<in> A \<times> B \<longleftrightarrow> fst x \<in> A \<and> snd x \<in> B"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5024
by (induct x) simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5025
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  5026
lemma seq_compact_Times: "seq_compact s \<Longrightarrow> seq_compact t \<Longrightarrow> seq_compact (s \<times> t)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  5027
unfolding seq_compact_def
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5028
apply clarify
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5029
apply (drule_tac x="fst \<circ> f" in spec)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5030
apply (drule mp, simp add: mem_Times_iff)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5031
apply (clarify, rename_tac l1 r1)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5032
apply (drule_tac x="snd \<circ> f \<circ> r1" in spec)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5033
apply (drule mp, simp add: mem_Times_iff)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5034
apply (clarify, rename_tac l2 r2)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5035
apply (rule_tac x="(l1, l2)" in rev_bexI, simp)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5036
apply (rule_tac x="r1 \<circ> r2" in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5037
apply (rule conjI, simp add: subseq_def)
50972
d2c6a0a7fcdf tuned proof
hoelzl
parents: 50971
diff changeset
  5038
apply (drule_tac f=r2 in LIMSEQ_subseq_LIMSEQ, assumption)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5039
apply (drule (1) tendsto_Pair) back
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5040
apply (simp add: o_def)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5041
done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5042
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  5043
text {* Generalize to @{class topological_space} *}
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  5044
lemma compact_Times: 
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  5045
  fixes s :: "'a::metric_space set" and t :: "'b::metric_space set"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  5046
  shows "compact s \<Longrightarrow> compact t \<Longrightarrow> compact (s \<times> t)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  5047
  unfolding compact_eq_seq_compact_metric by (rule seq_compact_Times)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  5048
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  5049
text{* Hence some useful properties follow quite easily. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5050
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5051
lemma compact_scaling:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5052
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5053
  assumes "compact s"  shows "compact ((\<lambda>x. c *\<^sub>R x) ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5054
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5055
  let ?f = "\<lambda>x. scaleR c x"
44282
f0de18b62d63 remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents: 44252
diff changeset
  5056
  have *:"bounded_linear ?f" by (rule bounded_linear_scaleR_right)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5057
  show ?thesis using compact_continuous_image[of s ?f] continuous_at_imp_continuous_on[of s ?f]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5058
    using linear_continuous_at[OF *] assms by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5059
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5060
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5061
lemma compact_negations:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5062
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5063
  assumes "compact s"  shows "compact ((\<lambda>x. -x) ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5064
  using compact_scaling [OF assms, of "- 1"] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5065
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5066
lemma compact_sums:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5067
  fixes s t :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5068
  assumes "compact s"  "compact t"  shows "compact {x + y | x y. x \<in> s \<and> y \<in> t}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5069
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5070
  have *:"{x + y | x y. x \<in> s \<and> y \<in> t} = (\<lambda>z. fst z + snd z) ` (s \<times> t)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5071
    apply auto unfolding image_iff apply(rule_tac x="(xa, y)" in bexI) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5072
  have "continuous_on (s \<times> t) (\<lambda>z. fst z + snd z)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5073
    unfolding continuous_on by (rule ballI) (intro tendsto_intros)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5074
  thus ?thesis unfolding * using compact_continuous_image compact_Times [OF assms] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5075
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5076
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5077
lemma compact_differences:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5078
  fixes s t :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5079
  assumes "compact s" "compact t"  shows "compact {x - y | x y. x \<in> s \<and> y \<in> t}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5080
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5081
  have "{x - y | x y. x\<in>s \<and> y \<in> t} =  {x + y | x y. x \<in> s \<and> y \<in> (uminus ` t)}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5082
    apply auto apply(rule_tac x= xa in exI) apply auto apply(rule_tac x=xa in exI) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5083
  thus ?thesis using compact_sums[OF assms(1) compact_negations[OF assms(2)]] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5084
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5085
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5086
lemma compact_translation:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5087
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5088
  assumes "compact s"  shows "compact ((\<lambda>x. a + x) ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5089
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5090
  have "{x + y |x y. x \<in> s \<and> y \<in> {a}} = (\<lambda>x. a + x) ` s" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5091
  thus ?thesis using compact_sums[OF assms compact_sing[of a]] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5092
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5093
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5094
lemma compact_affinity:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5095
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5096
  assumes "compact s"  shows "compact ((\<lambda>x. a + c *\<^sub>R x) ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5097
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5098
  have "op + a ` op *\<^sub>R c ` s = (\<lambda>x. a + c *\<^sub>R x) ` s" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5099
  thus ?thesis using compact_translation[OF compact_scaling[OF assms], of a c] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5100
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5101
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  5102
text {* Hence we get the following. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5103
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5104
lemma compact_sup_maxdistance:
50973
4a2c82644889 generalized diameter from real_normed_vector to metric_space
hoelzl
parents: 50972
diff changeset
  5105
  fixes s :: "'a::metric_space set"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5106
  assumes "compact s"  "s \<noteq> {}"
50973
4a2c82644889 generalized diameter from real_normed_vector to metric_space
hoelzl
parents: 50972
diff changeset
  5107
  shows "\<exists>x\<in>s. \<exists>y\<in>s. \<forall>u\<in>s. \<forall>v\<in>s. dist u v \<le> dist x y"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5108
proof-
50973
4a2c82644889 generalized diameter from real_normed_vector to metric_space
hoelzl
parents: 50972
diff changeset
  5109
  have "compact (s \<times> s)" using `compact s` by (intro compact_Times)
4a2c82644889 generalized diameter from real_normed_vector to metric_space
hoelzl
parents: 50972
diff changeset
  5110
  moreover have "s \<times> s \<noteq> {}" using `s \<noteq> {}` by auto
4a2c82644889 generalized diameter from real_normed_vector to metric_space
hoelzl
parents: 50972
diff changeset
  5111
  moreover have "continuous_on (s \<times> s) (\<lambda>x. dist (fst x) (snd x))"
4a2c82644889 generalized diameter from real_normed_vector to metric_space
hoelzl
parents: 50972
diff changeset
  5112
    by (intro continuous_at_imp_continuous_on ballI continuous_dist
4a2c82644889 generalized diameter from real_normed_vector to metric_space
hoelzl
parents: 50972
diff changeset
  5113
      continuous_isCont[THEN iffD1] isCont_fst isCont_snd isCont_ident)
4a2c82644889 generalized diameter from real_normed_vector to metric_space
hoelzl
parents: 50972
diff changeset
  5114
  ultimately show ?thesis
4a2c82644889 generalized diameter from real_normed_vector to metric_space
hoelzl
parents: 50972
diff changeset
  5115
    using continuous_attains_sup[of "s \<times> s" "\<lambda>x. dist (fst x) (snd x)"] by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5116
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5117
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  5118
text {* We can state this in terms of diameter of a set. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5119
50973
4a2c82644889 generalized diameter from real_normed_vector to metric_space
hoelzl
parents: 50972
diff changeset
  5120
definition "diameter s = (if s = {} then 0::real else Sup {dist x y | x y. x \<in> s \<and> y \<in> s})"
4a2c82644889 generalized diameter from real_normed_vector to metric_space
hoelzl
parents: 50972
diff changeset
  5121
4a2c82644889 generalized diameter from real_normed_vector to metric_space
hoelzl
parents: 50972
diff changeset
  5122
lemma diameter_bounded_bound:
4a2c82644889 generalized diameter from real_normed_vector to metric_space
hoelzl
parents: 50972
diff changeset
  5123
  fixes s :: "'a :: metric_space set"
4a2c82644889 generalized diameter from real_normed_vector to metric_space
hoelzl
parents: 50972
diff changeset
  5124
  assumes s: "bounded s" "x \<in> s" "y \<in> s"
4a2c82644889 generalized diameter from real_normed_vector to metric_space
hoelzl
parents: 50972
diff changeset
  5125
  shows "dist x y \<le> diameter s"
4a2c82644889 generalized diameter from real_normed_vector to metric_space
hoelzl
parents: 50972
diff changeset
  5126
proof -
4a2c82644889 generalized diameter from real_normed_vector to metric_space
hoelzl
parents: 50972
diff changeset
  5127
  let ?D = "{dist x y |x y. x \<in> s \<and> y \<in> s}"
4a2c82644889 generalized diameter from real_normed_vector to metric_space
hoelzl
parents: 50972
diff changeset
  5128
  from s obtain z d where z: "\<And>x. x \<in> s \<Longrightarrow> dist z x \<le> d"
4a2c82644889 generalized diameter from real_normed_vector to metric_space
hoelzl
parents: 50972
diff changeset
  5129
    unfolding bounded_def by auto
4a2c82644889 generalized diameter from real_normed_vector to metric_space
hoelzl
parents: 50972
diff changeset
  5130
  have "dist x y \<le> Sup ?D"
4a2c82644889 generalized diameter from real_normed_vector to metric_space
hoelzl
parents: 50972
diff changeset
  5131
  proof (rule Sup_upper, safe)
4a2c82644889 generalized diameter from real_normed_vector to metric_space
hoelzl
parents: 50972
diff changeset
  5132
    fix a b assume "a \<in> s" "b \<in> s"
4a2c82644889 generalized diameter from real_normed_vector to metric_space
hoelzl
parents: 50972
diff changeset
  5133
    with z[of a] z[of b] dist_triangle[of a b z]
4a2c82644889 generalized diameter from real_normed_vector to metric_space
hoelzl
parents: 50972
diff changeset
  5134
    show "dist a b \<le> 2 * d"
4a2c82644889 generalized diameter from real_normed_vector to metric_space
hoelzl
parents: 50972
diff changeset
  5135
      by (simp add: dist_commute)
4a2c82644889 generalized diameter from real_normed_vector to metric_space
hoelzl
parents: 50972
diff changeset
  5136
  qed (insert s, auto)
4a2c82644889 generalized diameter from real_normed_vector to metric_space
hoelzl
parents: 50972
diff changeset
  5137
  with `x \<in> s` show ?thesis
4a2c82644889 generalized diameter from real_normed_vector to metric_space
hoelzl
parents: 50972
diff changeset
  5138
    by (auto simp add: diameter_def)
4a2c82644889 generalized diameter from real_normed_vector to metric_space
hoelzl
parents: 50972
diff changeset
  5139
qed
4a2c82644889 generalized diameter from real_normed_vector to metric_space
hoelzl
parents: 50972
diff changeset
  5140
4a2c82644889 generalized diameter from real_normed_vector to metric_space
hoelzl
parents: 50972
diff changeset
  5141
lemma diameter_lower_bounded:
4a2c82644889 generalized diameter from real_normed_vector to metric_space
hoelzl
parents: 50972
diff changeset
  5142
  fixes s :: "'a :: metric_space set"
4a2c82644889 generalized diameter from real_normed_vector to metric_space
hoelzl
parents: 50972
diff changeset
  5143
  assumes s: "bounded s" and d: "0 < d" "d < diameter s"
4a2c82644889 generalized diameter from real_normed_vector to metric_space
hoelzl
parents: 50972
diff changeset
  5144
  shows "\<exists>x\<in>s. \<exists>y\<in>s. d < dist x y"
4a2c82644889 generalized diameter from real_normed_vector to metric_space
hoelzl
parents: 50972
diff changeset
  5145
proof (rule ccontr)
4a2c82644889 generalized diameter from real_normed_vector to metric_space
hoelzl
parents: 50972
diff changeset
  5146
  let ?D = "{dist x y |x y. x \<in> s \<and> y \<in> s}"
4a2c82644889 generalized diameter from real_normed_vector to metric_space
hoelzl
parents: 50972
diff changeset
  5147
  assume contr: "\<not> ?thesis"
4a2c82644889 generalized diameter from real_normed_vector to metric_space
hoelzl
parents: 50972
diff changeset
  5148
  moreover
4a2c82644889 generalized diameter from real_normed_vector to metric_space
hoelzl
parents: 50972
diff changeset
  5149
  from d have "s \<noteq> {}"
4a2c82644889 generalized diameter from real_normed_vector to metric_space
hoelzl
parents: 50972
diff changeset
  5150
    by (auto simp: diameter_def)
4a2c82644889 generalized diameter from real_normed_vector to metric_space
hoelzl
parents: 50972
diff changeset
  5151
  then have "?D \<noteq> {}" by auto
4a2c82644889 generalized diameter from real_normed_vector to metric_space
hoelzl
parents: 50972
diff changeset
  5152
  ultimately have "Sup ?D \<le> d"
4a2c82644889 generalized diameter from real_normed_vector to metric_space
hoelzl
parents: 50972
diff changeset
  5153
    by (intro Sup_least) (auto simp: not_less)
4a2c82644889 generalized diameter from real_normed_vector to metric_space
hoelzl
parents: 50972
diff changeset
  5154
  with `d < diameter s` `s \<noteq> {}` show False
4a2c82644889 generalized diameter from real_normed_vector to metric_space
hoelzl
parents: 50972
diff changeset
  5155
    by (auto simp: diameter_def)
4a2c82644889 generalized diameter from real_normed_vector to metric_space
hoelzl
parents: 50972
diff changeset
  5156
qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5157
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5158
lemma diameter_bounded:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5159
  assumes "bounded s"
50973
4a2c82644889 generalized diameter from real_normed_vector to metric_space
hoelzl
parents: 50972
diff changeset
  5160
  shows "\<forall>x\<in>s. \<forall>y\<in>s. dist x y \<le> diameter s"
4a2c82644889 generalized diameter from real_normed_vector to metric_space
hoelzl
parents: 50972
diff changeset
  5161
        "\<forall>d>0. d < diameter s \<longrightarrow> (\<exists>x\<in>s. \<exists>y\<in>s. dist x y > d)"
4a2c82644889 generalized diameter from real_normed_vector to metric_space
hoelzl
parents: 50972
diff changeset
  5162
  using diameter_bounded_bound[of s] diameter_lower_bounded[of s] assms
4a2c82644889 generalized diameter from real_normed_vector to metric_space
hoelzl
parents: 50972
diff changeset
  5163
  by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5164
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5165
lemma diameter_compact_attained:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5166
  assumes "compact s"  "s \<noteq> {}"
50973
4a2c82644889 generalized diameter from real_normed_vector to metric_space
hoelzl
parents: 50972
diff changeset
  5167
  shows "\<exists>x\<in>s. \<exists>y\<in>s. dist x y = diameter s"
4a2c82644889 generalized diameter from real_normed_vector to metric_space
hoelzl
parents: 50972
diff changeset
  5168
proof -
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5169
  have b:"bounded s" using assms(1) by (rule compact_imp_bounded)
50973
4a2c82644889 generalized diameter from real_normed_vector to metric_space
hoelzl
parents: 50972
diff changeset
  5170
  then obtain x y where xys:"x\<in>s" "y\<in>s" and xy:"\<forall>u\<in>s. \<forall>v\<in>s. dist u v \<le> dist x y"
4a2c82644889 generalized diameter from real_normed_vector to metric_space
hoelzl
parents: 50972
diff changeset
  5171
    using compact_sup_maxdistance[OF assms] by auto
4a2c82644889 generalized diameter from real_normed_vector to metric_space
hoelzl
parents: 50972
diff changeset
  5172
  hence "diameter s \<le> dist x y"
36362
06475a1547cb fix lots of looping simp calls and other warnings
huffman
parents: 36360
diff changeset
  5173
    unfolding diameter_def by clarsimp (rule Sup_least, fast+)
33324
51eb2ffa2189 Tidied up some very ugly proofs
paulson
parents: 33270
diff changeset
  5174
  thus ?thesis
36362
06475a1547cb fix lots of looping simp calls and other warnings
huffman
parents: 36360
diff changeset
  5175
    by (metis b diameter_bounded_bound order_antisym xys)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5176
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5177
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  5178
text {* Related results with closure as the conclusion. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5179
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5180
lemma closed_scaling:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5181
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5182
  assumes "closed s" shows "closed ((\<lambda>x. c *\<^sub>R x) ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5183
proof(cases "s={}")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5184
  case True thus ?thesis by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5185
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5186
  case False
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5187
  show ?thesis
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5188
  proof(cases "c=0")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5189
    have *:"(\<lambda>x. 0) ` s = {0}" using `s\<noteq>{}` by auto
36668
941ba2da372e simplify definition of t1_space; generalize lemma closed_sing and related lemmas
huffman
parents: 36667
diff changeset
  5190
    case True thus ?thesis apply auto unfolding * by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5191
  next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5192
    case False
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5193
    { fix x l assume as:"\<forall>n::nat. x n \<in> scaleR c ` s"  "(x ---> l) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5194
      { fix n::nat have "scaleR (1 / c) (x n) \<in> s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5195
          using as(1)[THEN spec[where x=n]]
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5196
          using `c\<noteq>0` by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5197
      }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5198
      moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5199
      { fix e::real assume "e>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5200
        hence "0 < e *\<bar>c\<bar>"  using `c\<noteq>0` mult_pos_pos[of e "abs c"] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5201
        then obtain N where "\<forall>n\<ge>N. dist (x n) l < e * \<bar>c\<bar>"
44907
93943da0a010 remove redundant lemma Lim_sequentially in favor of lemma LIMSEQ_def
huffman
parents: 44905
diff changeset
  5202
          using as(2)[unfolded LIMSEQ_def, THEN spec[where x="e * abs c"]] by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5203
        hence "\<exists>N. \<forall>n\<ge>N. dist (scaleR (1 / c) (x n)) (scaleR (1 / c) l) < e"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5204
          unfolding dist_norm unfolding scaleR_right_diff_distrib[THEN sym]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5205
          using mult_imp_div_pos_less[of "abs c" _ e] `c\<noteq>0` by auto  }
44907
93943da0a010 remove redundant lemma Lim_sequentially in favor of lemma LIMSEQ_def
huffman
parents: 44905
diff changeset
  5206
      hence "((\<lambda>n. scaleR (1 / c) (x n)) ---> scaleR (1 / c) l) sequentially" unfolding LIMSEQ_def by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5207
      ultimately have "l \<in> scaleR c ` s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5208
        using assms[unfolded closed_sequential_limits, THEN spec[where x="\<lambda>n. scaleR (1/c) (x n)"], THEN spec[where x="scaleR (1/c) l"]]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5209
        unfolding image_iff using `c\<noteq>0` apply(rule_tac x="scaleR (1 / c) l" in bexI) by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5210
    thus ?thesis unfolding closed_sequential_limits by fast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5211
  qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5212
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5213
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5214
lemma closed_negations:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5215
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5216
  assumes "closed s"  shows "closed ((\<lambda>x. -x) ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5217
  using closed_scaling[OF assms, of "- 1"] by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5218
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5219
lemma compact_closed_sums:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5220
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5221
  assumes "compact s"  "closed t"  shows "closed {x + y | x y. x \<in> s \<and> y \<in> t}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5222
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5223
  let ?S = "{x + y |x y. x \<in> s \<and> y \<in> t}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5224
  { fix x l assume as:"\<forall>n. x n \<in> ?S"  "(x ---> l) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5225
    from as(1) obtain f where f:"\<forall>n. x n = fst (f n) + snd (f n)"  "\<forall>n. fst (f n) \<in> s"  "\<forall>n. snd (f n) \<in> t"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5226
      using choice[of "\<lambda>n y. x n = (fst y) + (snd y) \<and> fst y \<in> s \<and> snd y \<in> t"] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5227
    obtain l' r where "l'\<in>s" and r:"subseq r" and lr:"(((\<lambda>n. fst (f n)) \<circ> r) ---> l') sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5228
      using assms(1)[unfolded compact_def, THEN spec[where x="\<lambda> n. fst (f n)"]] using f(2) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5229
    have "((\<lambda>n. snd (f (r n))) ---> l - l') sequentially"
50972
d2c6a0a7fcdf tuned proof
hoelzl
parents: 50971
diff changeset
  5230
      using tendsto_diff[OF LIMSEQ_subseq_LIMSEQ[OF as(2) r] lr] and f(1) unfolding o_def by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5231
    hence "l - l' \<in> t"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5232
      using assms(2)[unfolded closed_sequential_limits, THEN spec[where x="\<lambda> n. snd (f (r n))"], THEN spec[where x="l - l'"]]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5233
      using f(3) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5234
    hence "l \<in> ?S" using `l' \<in> s` apply auto apply(rule_tac x=l' in exI) apply(rule_tac x="l - l'" in exI) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5235
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5236
  thus ?thesis unfolding closed_sequential_limits by fast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5237
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5238
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5239
lemma closed_compact_sums:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5240
  fixes s t :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5241
  assumes "closed s"  "compact t"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5242
  shows "closed {x + y | x y. x \<in> s \<and> y \<in> t}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5243
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5244
  have "{x + y |x y. x \<in> t \<and> y \<in> s} = {x + y |x y. x \<in> s \<and> y \<in> t}" apply auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5245
    apply(rule_tac x=y in exI) apply auto apply(rule_tac x=y in exI) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5246
  thus ?thesis using compact_closed_sums[OF assms(2,1)] by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5247
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5248
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5249
lemma compact_closed_differences:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5250
  fixes s t :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5251
  assumes "compact s"  "closed t"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5252
  shows "closed {x - y | x y. x \<in> s \<and> y \<in> t}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5253
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5254
  have "{x + y |x y. x \<in> s \<and> y \<in> uminus ` t} =  {x - y |x y. x \<in> s \<and> y \<in> t}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5255
    apply auto apply(rule_tac x=xa in exI) apply auto apply(rule_tac x=xa in exI) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5256
  thus ?thesis using compact_closed_sums[OF assms(1) closed_negations[OF assms(2)]] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5257
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5258
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5259
lemma closed_compact_differences:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5260
  fixes s t :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5261
  assumes "closed s" "compact t"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5262
  shows "closed {x - y | x y. x \<in> s \<and> y \<in> t}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5263
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5264
  have "{x + y |x y. x \<in> s \<and> y \<in> uminus ` t} = {x - y |x y. x \<in> s \<and> y \<in> t}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5265
    apply auto apply(rule_tac x=xa in exI) apply auto apply(rule_tac x=xa in exI) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5266
 thus ?thesis using closed_compact_sums[OF assms(1) compact_negations[OF assms(2)]] by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5267
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5268
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5269
lemma closed_translation:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5270
  fixes a :: "'a::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5271
  assumes "closed s"  shows "closed ((\<lambda>x. a + x) ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5272
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5273
  have "{a + y |y. y \<in> s} = (op + a ` s)" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5274
  thus ?thesis using compact_closed_sums[OF compact_sing[of a] assms] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5275
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5276
34105
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
  5277
lemma translation_Compl:
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
  5278
  fixes a :: "'a::ab_group_add"
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
  5279
  shows "(\<lambda>x. a + x) ` (- t) = - ((\<lambda>x. a + x) ` t)"
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
  5280
  apply (auto simp add: image_iff) apply(rule_tac x="x - a" in bexI) by auto
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
  5281
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5282
lemma translation_UNIV:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5283
  fixes a :: "'a::ab_group_add" shows "range (\<lambda>x. a + x) = UNIV"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5284
  apply (auto simp add: image_iff) apply(rule_tac x="x - a" in exI) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5285
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5286
lemma translation_diff:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5287
  fixes a :: "'a::ab_group_add"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5288
  shows "(\<lambda>x. a + x) ` (s - t) = ((\<lambda>x. a + x) ` s) - ((\<lambda>x. a + x) ` t)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5289
  by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5290
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5291
lemma closure_translation:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5292
  fixes a :: "'a::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5293
  shows "closure ((\<lambda>x. a + x) ` s) = (\<lambda>x. a + x) ` (closure s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5294
proof-
34105
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
  5295
  have *:"op + a ` (- s) = - op + a ` s"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5296
    apply auto unfolding image_iff apply(rule_tac x="x - a" in bexI) by auto
34105
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
  5297
  show ?thesis unfolding closure_interior translation_Compl
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
  5298
    using interior_translation[of a "- s"] unfolding * by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5299
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5300
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5301
lemma frontier_translation:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5302
  fixes a :: "'a::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5303
  shows "frontier((\<lambda>x. a + x) ` s) = (\<lambda>x. a + x) ` (frontier s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5304
  unfolding frontier_def translation_diff interior_translation closure_translation by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5305
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  5306
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  5307
subsection {* Separation between points and sets *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5308
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5309
lemma separate_point_closed:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5310
  fixes s :: "'a::heine_borel set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5311
  shows "closed s \<Longrightarrow> a \<notin> s  ==> (\<exists>d>0. \<forall>x\<in>s. d \<le> dist a x)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5312
proof(cases "s = {}")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5313
  case True
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5314
  thus ?thesis by(auto intro!: exI[where x=1])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5315
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5316
  case False
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5317
  assume "closed s" "a \<notin> s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5318
  then obtain x where "x\<in>s" "\<forall>y\<in>s. dist a x \<le> dist a y" using `s \<noteq> {}` distance_attains_inf [of s a] by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5319
  with `x\<in>s` show ?thesis using dist_pos_lt[of a x] and`a \<notin> s` by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5320
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5321
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5322
lemma separate_compact_closed:
50949
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5323
  fixes s t :: "'a::heine_borel set"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5324
  assumes "compact s" and "closed t" and "s \<inter> t = {}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5325
  shows "\<exists>d>0. \<forall>x\<in>s. \<forall>y\<in>t. d \<le> dist x y"
50949
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5326
proof - (* FIXME: long proof *)
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5327
  let ?T = "\<Union>x\<in>s. { ball x (d / 2) | d. 0 < d \<and> (\<forall>y\<in>t. d \<le> dist x y) }"
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5328
  note `compact s`
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5329
  moreover have "\<forall>t\<in>?T. open t" by auto
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5330
  moreover have "s \<subseteq> \<Union>?T"
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5331
    apply auto
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5332
    apply (rule rev_bexI, assumption)
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5333
    apply (subgoal_tac "x \<notin> t")
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5334
    apply (drule separate_point_closed [OF `closed t`])
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5335
    apply clarify
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5336
    apply (rule_tac x="ball x (d / 2)" in exI)
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5337
    apply simp
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5338
    apply fast
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5339
    apply (cut_tac assms(3))
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5340
    apply auto
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5341
    done
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5342
  ultimately obtain U where "U \<subseteq> ?T" and "finite U" and "s \<subseteq> \<Union>U"
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5343
    by (rule compactE)
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5344
  from `finite U` and `U \<subseteq> ?T` have "\<exists>d>0. \<forall>x\<in>\<Union>U. \<forall>y\<in>t. d \<le> dist x y"
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5345
    apply (induct set: finite)
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5346
    apply simp
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5347
    apply (rule exI)
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5348
    apply (rule zero_less_one)
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5349
    apply clarsimp
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5350
    apply (rename_tac y e)
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5351
    apply (rule_tac x="min d (e / 2)" in exI)
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5352
    apply simp
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5353
    apply (subst ball_Un)
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5354
    apply (rule conjI)
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5355
    apply (intro ballI, rename_tac z)
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5356
    apply (rule min_max.le_infI2)
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5357
    apply (simp only: mem_ball)
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5358
    apply (drule (1) bspec)
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5359
    apply (cut_tac x=y and y=x and z=z in dist_triangle, arith)
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5360
    apply simp
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5361
    apply (intro ballI)
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5362
    apply (rule min_max.le_infI1)
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5363
    apply simp
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5364
    done
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5365
  with `s \<subseteq> \<Union>U` show ?thesis
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5366
    by fast
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5367
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5368
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5369
lemma separate_closed_compact:
50949
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5370
  fixes s t :: "'a::heine_borel set"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5371
  assumes "closed s" and "compact t" and "s \<inter> t = {}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5372
  shows "\<exists>d>0. \<forall>x\<in>s. \<forall>y\<in>t. d \<le> dist x y"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5373
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5374
  have *:"t \<inter> s = {}" using assms(3) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5375
  show ?thesis using separate_compact_closed[OF assms(2,1) *]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5376
    apply auto apply(rule_tac x=d in exI) apply auto apply (erule_tac x=y in ballE)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5377
    by (auto simp add: dist_commute)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5378
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5379
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  5380
36439
a65320184de9 move intervals section heading
huffman
parents: 36438
diff changeset
  5381
subsection {* Intervals *}
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5382
  
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5383
lemma interval: fixes a :: "'a::ordered_euclidean_space" shows
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5384
  "{a <..< b} = {x::'a. \<forall>i\<in>Basis. a\<bullet>i < x\<bullet>i \<and> x\<bullet>i < b\<bullet>i}" and
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5385
  "{a .. b} = {x::'a. \<forall>i\<in>Basis. a\<bullet>i \<le> x\<bullet>i \<and> x\<bullet>i \<le> b\<bullet>i}"
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
  5386
  by(auto simp add:set_eq_iff eucl_le[where 'a='a] eucl_less[where 'a='a])
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5387
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5388
lemma mem_interval: fixes a :: "'a::ordered_euclidean_space" shows
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5389
  "x \<in> {a<..<b} \<longleftrightarrow> (\<forall>i\<in>Basis. a\<bullet>i < x\<bullet>i \<and> x\<bullet>i < b\<bullet>i)"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5390
  "x \<in> {a .. b} \<longleftrightarrow> (\<forall>i\<in>Basis. a\<bullet>i \<le> x\<bullet>i \<and> x\<bullet>i \<le> b\<bullet>i)"
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
  5391
  using interval[of a b] by(auto simp add: set_eq_iff eucl_le[where 'a='a] eucl_less[where 'a='a])
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5392
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5393
lemma interval_eq_empty: fixes a :: "'a::ordered_euclidean_space" shows
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5394
 "({a <..< b} = {} \<longleftrightarrow> (\<exists>i\<in>Basis. b\<bullet>i \<le> a\<bullet>i))" (is ?th1) and
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5395
 "({a  ..  b} = {} \<longleftrightarrow> (\<exists>i\<in>Basis. b\<bullet>i < a\<bullet>i))" (is ?th2)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5396
proof-
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5397
  { fix i x assume i:"i\<in>Basis" and as:"b\<bullet>i \<le> a\<bullet>i" and x:"x\<in>{a <..< b}"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5398
    hence "a \<bullet> i < x \<bullet> i \<and> x \<bullet> i < b \<bullet> i" unfolding mem_interval by auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5399
    hence "a\<bullet>i < b\<bullet>i" by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5400
    hence False using as by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5401
  moreover
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5402
  { assume as:"\<forall>i\<in>Basis. \<not> (b\<bullet>i \<le> a\<bullet>i)"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5403
    let ?x = "(1/2) *\<^sub>R (a + b)"
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5404
    { fix i :: 'a assume i:"i\<in>Basis" 
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5405
      have "a\<bullet>i < b\<bullet>i" using as[THEN bspec[where x=i]] i by auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5406
      hence "a\<bullet>i < ((1/2) *\<^sub>R (a+b)) \<bullet> i" "((1/2) *\<^sub>R (a+b)) \<bullet> i < b\<bullet>i"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5407
        by (auto simp: inner_add_left) }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5408
    hence "{a <..< b} \<noteq> {}" using mem_interval(1)[of "?x" a b] by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5409
  ultimately show ?th1 by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5410
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5411
  { fix i x assume i:"i\<in>Basis" and as:"b\<bullet>i < a\<bullet>i" and x:"x\<in>{a .. b}"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5412
    hence "a \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> b \<bullet> i" unfolding mem_interval by auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5413
    hence "a\<bullet>i \<le> b\<bullet>i" by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5414
    hence False using as by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5415
  moreover
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5416
  { assume as:"\<forall>i\<in>Basis. \<not> (b\<bullet>i < a\<bullet>i)"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5417
    let ?x = "(1/2) *\<^sub>R (a + b)"
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5418
    { fix i :: 'a assume i:"i\<in>Basis"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5419
      have "a\<bullet>i \<le> b\<bullet>i" using as[THEN bspec[where x=i]] i by auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5420
      hence "a\<bullet>i \<le> ((1/2) *\<^sub>R (a+b)) \<bullet> i" "((1/2) *\<^sub>R (a+b)) \<bullet> i \<le> b\<bullet>i"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5421
        by (auto simp: inner_add_left) }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5422
    hence "{a .. b} \<noteq> {}" using mem_interval(2)[of "?x" a b] by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5423
  ultimately show ?th2 by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5424
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5425
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5426
lemma interval_ne_empty: fixes a :: "'a::ordered_euclidean_space" shows
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5427
  "{a  ..  b} \<noteq> {} \<longleftrightarrow> (\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i)" and
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5428
  "{a <..< b} \<noteq> {} \<longleftrightarrow> (\<forall>i\<in>Basis. a\<bullet>i < b\<bullet>i)"
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 44668
diff changeset
  5429
  unfolding interval_eq_empty[of a b] by fastforce+
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5430
44584
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  5431
lemma interval_sing:
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  5432
  fixes a :: "'a::ordered_euclidean_space"
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  5433
  shows "{a .. a} = {a}" and "{a<..<a} = {}"
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  5434
  unfolding set_eq_iff mem_interval eq_iff [symmetric]
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5435
  by (auto intro: euclidean_eqI simp: ex_in_conv)
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5436
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5437
lemma subset_interval_imp: fixes a :: "'a::ordered_euclidean_space" shows
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5438
 "(\<forall>i\<in>Basis. a\<bullet>i \<le> c\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i) \<Longrightarrow> {c .. d} \<subseteq> {a .. b}" and
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5439
 "(\<forall>i\<in>Basis. a\<bullet>i < c\<bullet>i \<and> d\<bullet>i < b\<bullet>i) \<Longrightarrow> {c .. d} \<subseteq> {a<..<b}" and
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5440
 "(\<forall>i\<in>Basis. a\<bullet>i \<le> c\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i) \<Longrightarrow> {c<..<d} \<subseteq> {a .. b}" and
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5441
 "(\<forall>i\<in>Basis. a\<bullet>i \<le> c\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i) \<Longrightarrow> {c<..<d} \<subseteq> {a<..<b}"
44584
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  5442
  unfolding subset_eq[unfolded Ball_def] unfolding mem_interval
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  5443
  by (best intro: order_trans less_le_trans le_less_trans less_imp_le)+
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  5444
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  5445
lemma interval_open_subset_closed:
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  5446
  fixes a :: "'a::ordered_euclidean_space"
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  5447
  shows "{a<..<b} \<subseteq> {a .. b}"
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  5448
  unfolding subset_eq [unfolded Ball_def] mem_interval
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  5449
  by (fast intro: less_imp_le)
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5450
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5451
lemma subset_interval: fixes a :: "'a::ordered_euclidean_space" shows
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5452
 "{c .. d} \<subseteq> {a .. b} \<longleftrightarrow> (\<forall>i\<in>Basis. c\<bullet>i \<le> d\<bullet>i) --> (\<forall>i\<in>Basis. a\<bullet>i \<le> c\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i)" (is ?th1) and
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5453
 "{c .. d} \<subseteq> {a<..<b} \<longleftrightarrow> (\<forall>i\<in>Basis. c\<bullet>i \<le> d\<bullet>i) --> (\<forall>i\<in>Basis. a\<bullet>i < c\<bullet>i \<and> d\<bullet>i < b\<bullet>i)" (is ?th2) and
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5454
 "{c<..<d} \<subseteq> {a .. b} \<longleftrightarrow> (\<forall>i\<in>Basis. c\<bullet>i < d\<bullet>i) --> (\<forall>i\<in>Basis. a\<bullet>i \<le> c\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i)" (is ?th3) and
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5455
 "{c<..<d} \<subseteq> {a<..<b} \<longleftrightarrow> (\<forall>i\<in>Basis. c\<bullet>i < d\<bullet>i) --> (\<forall>i\<in>Basis. a\<bullet>i \<le> c\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i)" (is ?th4)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5456
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5457
  show ?th1 unfolding subset_eq and Ball_def and mem_interval by (auto intro: order_trans)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5458
  show ?th2 unfolding subset_eq and Ball_def and mem_interval by (auto intro: le_less_trans less_le_trans order_trans less_imp_le)
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5459
  { assume as: "{c<..<d} \<subseteq> {a .. b}" "\<forall>i\<in>Basis. c\<bullet>i < d\<bullet>i"
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5460
    hence "{c<..<d} \<noteq> {}" unfolding interval_eq_empty by auto
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5461
    fix i :: 'a assume i:"i\<in>Basis"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5462
    (** TODO combine the following two parts as done in the HOL_light version. **)
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5463
    { let ?x = "(\<Sum>j\<in>Basis. (if j=i then ((min (a\<bullet>j) (d\<bullet>j))+c\<bullet>j)/2 else (c\<bullet>j+d\<bullet>j)/2) *\<^sub>R j)::'a"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5464
      assume as2: "a\<bullet>i > c\<bullet>i"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5465
      { fix j :: 'a assume j:"j\<in>Basis"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5466
        hence "c \<bullet> j < ?x \<bullet> j \<and> ?x \<bullet> j < d \<bullet> j"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5467
          apply(cases "j=i") using as(2)[THEN bspec[where x=j]] i
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5468
          by (auto simp add: as2)  }
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5469
      hence "?x\<in>{c<..<d}" using i unfolding mem_interval by auto
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5470
      moreover
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5471
      have "?x\<notin>{a .. b}"
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5472
        unfolding mem_interval apply auto apply(rule_tac x=i in bexI)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5473
        using as(2)[THEN bspec[where x=i]] and as2 i
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5474
        by auto
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5475
      ultimately have False using as by auto  }
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5476
    hence "a\<bullet>i \<le> c\<bullet>i" by(rule ccontr)auto
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5477
    moreover
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5478
    { let ?x = "(\<Sum>j\<in>Basis. (if j=i then ((max (b\<bullet>j) (c\<bullet>j))+d\<bullet>j)/2 else (c\<bullet>j+d\<bullet>j)/2) *\<^sub>R j)::'a"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5479
      assume as2: "b\<bullet>i < d\<bullet>i"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5480
      { fix j :: 'a assume "j\<in>Basis"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5481
        hence "d \<bullet> j > ?x \<bullet> j \<and> ?x \<bullet> j > c \<bullet> j" 
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5482
          apply(cases "j=i") using as(2)[THEN bspec[where x=j]]
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5483
          by (auto simp add: as2) }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5484
      hence "?x\<in>{c<..<d}" unfolding mem_interval by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5485
      moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5486
      have "?x\<notin>{a .. b}"
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5487
        unfolding mem_interval apply auto apply(rule_tac x=i in bexI)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5488
        using as(2)[THEN bspec[where x=i]] and as2 using i
36362
06475a1547cb fix lots of looping simp calls and other warnings
huffman
parents: 36360
diff changeset
  5489
        by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5490
      ultimately have False using as by auto  }
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5491
    hence "b\<bullet>i \<ge> d\<bullet>i" by(rule ccontr)auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5492
    ultimately
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5493
    have "a\<bullet>i \<le> c\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i" by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5494
  } note part1 = this
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5495
  show ?th3
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5496
    unfolding subset_eq and Ball_def and mem_interval 
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5497
    apply(rule,rule,rule,rule) 
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5498
    apply(rule part1)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5499
    unfolding subset_eq and Ball_def and mem_interval
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5500
    prefer 4
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5501
    apply auto 
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5502
    by(erule_tac x=xa in allE,erule_tac x=xa in allE,fastforce)+ 
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5503
  { assume as:"{c<..<d} \<subseteq> {a<..<b}" "\<forall>i\<in>Basis. c\<bullet>i < d\<bullet>i"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5504
    fix i :: 'a assume i:"i\<in>Basis"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5505
    from as(1) have "{c<..<d} \<subseteq> {a..b}" using interval_open_subset_closed[of a b] by auto
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5506
    hence "a\<bullet>i \<le> c\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i" using part1 and as(2) using i by auto  } note * = this
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5507
  show ?th4 unfolding subset_eq and Ball_def and mem_interval 
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5508
    apply(rule,rule,rule,rule) apply(rule *) unfolding subset_eq and Ball_def and mem_interval prefer 4
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5509
    apply auto by(erule_tac x=xa in allE, simp)+ 
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5510
qed
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5511
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5512
lemma inter_interval: fixes a :: "'a::ordered_euclidean_space" shows
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5513
 "{a .. b} \<inter> {c .. d} =  {(\<Sum>i\<in>Basis. max (a\<bullet>i) (c\<bullet>i) *\<^sub>R i) .. (\<Sum>i\<in>Basis. min (b\<bullet>i) (d\<bullet>i) *\<^sub>R i)}"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5514
  unfolding set_eq_iff and Int_iff and mem_interval by auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5515
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5516
lemma disjoint_interval: fixes a::"'a::ordered_euclidean_space" shows
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5517
  "{a .. b} \<inter> {c .. d} = {} \<longleftrightarrow> (\<exists>i\<in>Basis. (b\<bullet>i < a\<bullet>i \<or> d\<bullet>i < c\<bullet>i \<or> b\<bullet>i < c\<bullet>i \<or> d\<bullet>i < a\<bullet>i))" (is ?th1) and
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5518
  "{a .. b} \<inter> {c<..<d} = {} \<longleftrightarrow> (\<exists>i\<in>Basis. (b\<bullet>i < a\<bullet>i \<or> d\<bullet>i \<le> c\<bullet>i \<or> b\<bullet>i \<le> c\<bullet>i \<or> d\<bullet>i \<le> a\<bullet>i))" (is ?th2) and
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5519
  "{a<..<b} \<inter> {c .. d} = {} \<longleftrightarrow> (\<exists>i\<in>Basis. (b\<bullet>i \<le> a\<bullet>i \<or> d\<bullet>i < c\<bullet>i \<or> b\<bullet>i \<le> c\<bullet>i \<or> d\<bullet>i \<le> a\<bullet>i))" (is ?th3) and
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5520
  "{a<..<b} \<inter> {c<..<d} = {} \<longleftrightarrow> (\<exists>i\<in>Basis. (b\<bullet>i \<le> a\<bullet>i \<or> d\<bullet>i \<le> c\<bullet>i \<or> b\<bullet>i \<le> c\<bullet>i \<or> d\<bullet>i \<le> a\<bullet>i))" (is ?th4)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5521
proof-
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5522
  let ?z = "(\<Sum>i\<in>Basis. (((max (a\<bullet>i) (c\<bullet>i)) + (min (b\<bullet>i) (d\<bullet>i))) / 2) *\<^sub>R i)::'a"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5523
  have **: "\<And>P Q. (\<And>i :: 'a. i \<in> Basis \<Longrightarrow> Q ?z i \<Longrightarrow> P i) \<Longrightarrow>
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5524
      (\<And>i x :: 'a. i \<in> Basis \<Longrightarrow> P i \<Longrightarrow> Q x i) \<Longrightarrow> (\<forall>x. \<exists>i\<in>Basis. Q x i) \<longleftrightarrow> (\<exists>i\<in>Basis. P i)" 
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5525
    by blast
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5526
  note * = set_eq_iff Int_iff empty_iff mem_interval ball_conj_distrib[symmetric] eq_False ball_simps(10)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5527
  show ?th1 unfolding * by (intro **) auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5528
  show ?th2 unfolding * by (intro **) auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5529
  show ?th3 unfolding * by (intro **) auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5530
  show ?th4 unfolding * by (intro **) auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5531
qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5532
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5533
(* Moved interval_open_subset_closed a bit upwards *)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5534
44250
9133bc634d9c simplify proofs of lemmas open_interval, closed_interval
huffman
parents: 44233
diff changeset
  5535
lemma open_interval[intro]:
9133bc634d9c simplify proofs of lemmas open_interval, closed_interval
huffman
parents: 44233
diff changeset
  5536
  fixes a b :: "'a::ordered_euclidean_space" shows "open {a<..<b}"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5537
proof-
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5538
  have "open (\<Inter>i\<in>Basis. (\<lambda>x. x\<bullet>i) -` {a\<bullet>i<..<b\<bullet>i})"
44250
9133bc634d9c simplify proofs of lemmas open_interval, closed_interval
huffman
parents: 44233
diff changeset
  5539
    by (intro open_INT finite_lessThan ballI continuous_open_vimage allI
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5540
      linear_continuous_at open_real_greaterThanLessThan finite_Basis bounded_linear_inner_left)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5541
  also have "(\<Inter>i\<in>Basis. (\<lambda>x. x\<bullet>i) -` {a\<bullet>i<..<b\<bullet>i}) = {a<..<b}"
44250
9133bc634d9c simplify proofs of lemmas open_interval, closed_interval
huffman
parents: 44233
diff changeset
  5542
    by (auto simp add: eucl_less [where 'a='a])
9133bc634d9c simplify proofs of lemmas open_interval, closed_interval
huffman
parents: 44233
diff changeset
  5543
  finally show "open {a<..<b}" .
9133bc634d9c simplify proofs of lemmas open_interval, closed_interval
huffman
parents: 44233
diff changeset
  5544
qed
9133bc634d9c simplify proofs of lemmas open_interval, closed_interval
huffman
parents: 44233
diff changeset
  5545
9133bc634d9c simplify proofs of lemmas open_interval, closed_interval
huffman
parents: 44233
diff changeset
  5546
lemma closed_interval[intro]:
9133bc634d9c simplify proofs of lemmas open_interval, closed_interval
huffman
parents: 44233
diff changeset
  5547
  fixes a b :: "'a::ordered_euclidean_space" shows "closed {a .. b}"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5548
proof-
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5549
  have "closed (\<Inter>i\<in>Basis. (\<lambda>x. x\<bullet>i) -` {a\<bullet>i .. b\<bullet>i})"
44250
9133bc634d9c simplify proofs of lemmas open_interval, closed_interval
huffman
parents: 44233
diff changeset
  5550
    by (intro closed_INT ballI continuous_closed_vimage allI
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5551
      linear_continuous_at closed_real_atLeastAtMost finite_Basis bounded_linear_inner_left)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5552
  also have "(\<Inter>i\<in>Basis. (\<lambda>x. x\<bullet>i) -` {a\<bullet>i .. b\<bullet>i}) = {a .. b}"
44250
9133bc634d9c simplify proofs of lemmas open_interval, closed_interval
huffman
parents: 44233
diff changeset
  5553
    by (auto simp add: eucl_le [where 'a='a])
9133bc634d9c simplify proofs of lemmas open_interval, closed_interval
huffman
parents: 44233
diff changeset
  5554
  finally show "closed {a .. b}" .
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5555
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5556
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  5557
lemma interior_closed_interval [intro]:
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  5558
  fixes a b :: "'a::ordered_euclidean_space"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  5559
  shows "interior {a..b} = {a<..<b}" (is "?L = ?R")
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5560
proof(rule subset_antisym)
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  5561
  show "?R \<subseteq> ?L" using interval_open_subset_closed open_interval
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  5562
    by (rule interior_maximal)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5563
next
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  5564
  { fix x assume "x \<in> interior {a..b}"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  5565
    then obtain s where s:"open s" "x \<in> s" "s \<subseteq> {a..b}" ..
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5566
    then obtain e where "e>0" and e:"\<forall>x'. dist x' x < e \<longrightarrow> x' \<in> {a..b}" unfolding open_dist and subset_eq by auto
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5567
    { fix i :: 'a assume i:"i\<in>Basis"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5568
      have "dist (x - (e / 2) *\<^sub>R i) x < e"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5569
           "dist (x + (e / 2) *\<^sub>R i) x < e"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5570
        unfolding dist_norm apply auto
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5571
        unfolding norm_minus_cancel using norm_Basis[OF i] `e>0` by auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5572
      hence "a \<bullet> i \<le> (x - (e / 2) *\<^sub>R i) \<bullet> i"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5573
                     "(x + (e / 2) *\<^sub>R i) \<bullet> i \<le> b \<bullet> i"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5574
        using e[THEN spec[where x="x - (e/2) *\<^sub>R i"]]
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5575
        and   e[THEN spec[where x="x + (e/2) *\<^sub>R i"]]
44584
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  5576
        unfolding mem_interval using i by blast+
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5577
      hence "a \<bullet> i < x \<bullet> i" and "x \<bullet> i < b \<bullet> i"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5578
        using `e>0` i by (auto simp: inner_diff_left inner_Basis inner_add_left) }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5579
    hence "x \<in> {a<..<b}" unfolding mem_interval by auto  }
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  5580
  thus "?L \<subseteq> ?R" ..
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5581
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5582
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5583
lemma bounded_closed_interval: fixes a :: "'a::ordered_euclidean_space" shows "bounded {a .. b}"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5584
proof-
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5585
  let ?b = "\<Sum>i\<in>Basis. \<bar>a\<bullet>i\<bar> + \<bar>b\<bullet>i\<bar>"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5586
  { fix x::"'a" assume x:"\<forall>i\<in>Basis. a \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> b \<bullet> i"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5587
    { fix i :: 'a assume "i\<in>Basis"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5588
      hence "\<bar>x\<bullet>i\<bar> \<le> \<bar>a\<bullet>i\<bar> + \<bar>b\<bullet>i\<bar>" using x[THEN bspec[where x=i]] by auto  }
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5589
    hence "(\<Sum>i\<in>Basis. \<bar>x \<bullet> i\<bar>) \<le> ?b" apply-apply(rule setsum_mono) by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5590
    hence "norm x \<le> ?b" using norm_le_l1[of x] by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5591
  thus ?thesis unfolding interval and bounded_iff by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5592
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5593
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5594
lemma bounded_interval: fixes a :: "'a::ordered_euclidean_space" shows
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5595
 "bounded {a .. b} \<and> bounded {a<..<b}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5596
  using bounded_closed_interval[of a b]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5597
  using interval_open_subset_closed[of a b]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5598
  using bounded_subset[of "{a..b}" "{a<..<b}"]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5599
  by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5600
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5601
lemma not_interval_univ: fixes a :: "'a::ordered_euclidean_space" shows
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5602
 "({a .. b} \<noteq> UNIV) \<and> ({a<..<b} \<noteq> UNIV)"
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5603
  using bounded_interval[of a b] by auto
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5604
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5605
lemma compact_interval: fixes a :: "'a::ordered_euclidean_space" shows "compact {a .. b}"
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  5606
  using bounded_closed_imp_seq_compact[of "{a..b}"] using bounded_interval[of a b]
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  5607
  by (auto simp: compact_eq_seq_compact_metric)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5608
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5609
lemma open_interval_midpoint: fixes a :: "'a::ordered_euclidean_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5610
  assumes "{a<..<b} \<noteq> {}" shows "((1/2) *\<^sub>R (a + b)) \<in> {a<..<b}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5611
proof-
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5612
  { fix i :: 'a assume "i\<in>Basis"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5613
    hence "a \<bullet> i < ((1 / 2) *\<^sub>R (a + b)) \<bullet> i \<and> ((1 / 2) *\<^sub>R (a + b)) \<bullet> i < b \<bullet> i"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5614
      using assms[unfolded interval_ne_empty, THEN bspec[where x=i]] by (auto simp: inner_add_left)  }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5615
  thus ?thesis unfolding mem_interval by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5616
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5617
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5618
lemma open_closed_interval_convex: fixes x :: "'a::ordered_euclidean_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5619
  assumes x:"x \<in> {a<..<b}" and y:"y \<in> {a .. b}" and e:"0 < e" "e \<le> 1"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5620
  shows "(e *\<^sub>R x + (1 - e) *\<^sub>R y) \<in> {a<..<b}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5621
proof-
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5622
  { fix i :: 'a assume i:"i\<in>Basis"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5623
    have "a \<bullet> i = e * (a \<bullet> i) + (1 - e) * (a \<bullet> i)" unfolding left_diff_distrib by simp
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5624
    also have "\<dots> < e * (x \<bullet> i) + (1 - e) * (y \<bullet> i)" apply(rule add_less_le_mono)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5625
      using e unfolding mult_less_cancel_left and mult_le_cancel_left apply simp_all
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5626
      using x unfolding mem_interval using i apply simp
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5627
      using y unfolding mem_interval using i apply simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5628
      done
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5629
    finally have "a \<bullet> i < (e *\<^sub>R x + (1 - e) *\<^sub>R y) \<bullet> i" unfolding inner_simps by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5630
    moreover {
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5631
    have "b \<bullet> i = e * (b\<bullet>i) + (1 - e) * (b\<bullet>i)" unfolding left_diff_distrib by simp
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5632
    also have "\<dots> > e * (x \<bullet> i) + (1 - e) * (y \<bullet> i)" apply(rule add_less_le_mono)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5633
      using e unfolding mult_less_cancel_left and mult_le_cancel_left apply simp_all
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5634
      using x unfolding mem_interval using i apply simp
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5635
      using y unfolding mem_interval using i apply simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5636
      done
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5637
    finally have "(e *\<^sub>R x + (1 - e) *\<^sub>R y) \<bullet> i < b \<bullet> i" unfolding inner_simps by auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5638
    } ultimately have "a \<bullet> i < (e *\<^sub>R x + (1 - e) *\<^sub>R y) \<bullet> i \<and> (e *\<^sub>R x + (1 - e) *\<^sub>R y) \<bullet> i < b \<bullet> i" by auto }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5639
  thus ?thesis unfolding mem_interval by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5640
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5641
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5642
lemma closure_open_interval: fixes a :: "'a::ordered_euclidean_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5643
  assumes "{a<..<b} \<noteq> {}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5644
  shows "closure {a<..<b} = {a .. b}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5645
proof-
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5646
  have ab:"a < b" using assms[unfolded interval_ne_empty] apply(subst eucl_less) by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5647
  let ?c = "(1 / 2) *\<^sub>R (a + b)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5648
  { fix x assume as:"x \<in> {a .. b}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5649
    def f == "\<lambda>n::nat. x + (inverse (real n + 1)) *\<^sub>R (?c - x)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5650
    { fix n assume fn:"f n < b \<longrightarrow> a < f n \<longrightarrow> f n = x" and xc:"x \<noteq> ?c"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5651
      have *:"0 < inverse (real n + 1)" "inverse (real n + 1) \<le> 1" unfolding inverse_le_1_iff by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5652
      have "(inverse (real n + 1)) *\<^sub>R ((1 / 2) *\<^sub>R (a + b)) + (1 - inverse (real n + 1)) *\<^sub>R x =
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5653
        x + (inverse (real n + 1)) *\<^sub>R (((1 / 2) *\<^sub>R (a + b)) - x)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5654
        by (auto simp add: algebra_simps)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5655
      hence "f n < b" and "a < f n" using open_closed_interval_convex[OF open_interval_midpoint[OF assms] as *] unfolding f_def by auto
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5656
      hence False using fn unfolding f_def using xc by auto  }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5657
    moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5658
    { assume "\<not> (f ---> x) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5659
      { fix e::real assume "e>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5660
        hence "\<exists>N::nat. inverse (real (N + 1)) < e" using real_arch_inv[of e] apply (auto simp add: Suc_pred') apply(rule_tac x="n - 1" in exI) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5661
        then obtain N::nat where "inverse (real (N + 1)) < e" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5662
        hence "\<forall>n\<ge>N. inverse (real n + 1) < e" by (auto, metis Suc_le_mono le_SucE less_imp_inverse_less nat_le_real_less order_less_trans real_of_nat_Suc real_of_nat_Suc_gt_zero)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5663
        hence "\<exists>N::nat. \<forall>n\<ge>N. inverse (real n + 1) < e" by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5664
      hence "((\<lambda>n. inverse (real n + 1)) ---> 0) sequentially"
44907
93943da0a010 remove redundant lemma Lim_sequentially in favor of lemma LIMSEQ_def
huffman
parents: 44905
diff changeset
  5665
        unfolding LIMSEQ_def by(auto simp add: dist_norm)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5666
      hence "(f ---> x) sequentially" unfolding f_def
44125
230a8665c919 mark some redundant theorems as legacy
huffman
parents: 44122
diff changeset
  5667
        using tendsto_add[OF tendsto_const, of "\<lambda>n::nat. (inverse (real n + 1)) *\<^sub>R ((1 / 2) *\<^sub>R (a + b) - x)" 0 sequentially x]
44282
f0de18b62d63 remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents: 44252
diff changeset
  5668
        using tendsto_scaleR [OF _ tendsto_const, of "\<lambda>n::nat. inverse (real n + 1)" 0 sequentially "((1 / 2) *\<^sub>R (a + b) - x)"] by auto  }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5669
    ultimately have "x \<in> closure {a<..<b}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5670
      using as and open_interval_midpoint[OF assms] unfolding closure_def unfolding islimpt_sequential by(cases "x=?c")auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5671
  thus ?thesis using closure_minimal[OF interval_open_subset_closed closed_interval, of a b] by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5672
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5673
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5674
lemma bounded_subset_open_interval_symmetric: fixes s::"('a::ordered_euclidean_space) set"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5675
  assumes "bounded s"  shows "\<exists>a. s \<subseteq> {-a<..<a}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5676
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5677
  obtain b where "b>0" and b:"\<forall>x\<in>s. norm x \<le> b" using assms[unfolded bounded_pos] by auto
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5678
  def a \<equiv> "(\<Sum>i\<in>Basis. (b + 1) *\<^sub>R i)::'a"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5679
  { fix x assume "x\<in>s"
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5680
    fix i :: 'a assume i:"i\<in>Basis"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5681
    hence "(-a)\<bullet>i < x\<bullet>i" and "x\<bullet>i < a\<bullet>i" using b[THEN bspec[where x=x], OF `x\<in>s`]
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5682
      and Basis_le_norm[OF i, of x] unfolding inner_simps and a_def by auto }
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5683
  thus ?thesis by(auto intro: exI[where x=a] simp add: eucl_less[where 'a='a])
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5684
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5685
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5686
lemma bounded_subset_open_interval:
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5687
  fixes s :: "('a::ordered_euclidean_space) set"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5688
  shows "bounded s ==> (\<exists>a b. s \<subseteq> {a<..<b})"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5689
  by (auto dest!: bounded_subset_open_interval_symmetric)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5690
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5691
lemma bounded_subset_closed_interval_symmetric:
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5692
  fixes s :: "('a::ordered_euclidean_space) set"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5693
  assumes "bounded s" shows "\<exists>a. s \<subseteq> {-a .. a}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5694
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5695
  obtain a where "s \<subseteq> {- a<..<a}" using bounded_subset_open_interval_symmetric[OF assms] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5696
  thus ?thesis using interval_open_subset_closed[of "-a" a] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5697
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5698
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5699
lemma bounded_subset_closed_interval:
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5700
  fixes s :: "('a::ordered_euclidean_space) set"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5701
  shows "bounded s ==> (\<exists>a b. s \<subseteq> {a .. b})"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5702
  using bounded_subset_closed_interval_symmetric[of s] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5703
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5704
lemma frontier_closed_interval:
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5705
  fixes a b :: "'a::ordered_euclidean_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5706
  shows "frontier {a .. b} = {a .. b} - {a<..<b}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5707
  unfolding frontier_def unfolding interior_closed_interval and closure_closed[OF closed_interval] ..
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5708
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5709
lemma frontier_open_interval:
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5710
  fixes a b :: "'a::ordered_euclidean_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5711
  shows "frontier {a<..<b} = (if {a<..<b} = {} then {} else {a .. b} - {a<..<b})"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5712
proof(cases "{a<..<b} = {}")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5713
  case True thus ?thesis using frontier_empty by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5714
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5715
  case False thus ?thesis unfolding frontier_def and closure_open_interval[OF False] and interior_open[OF open_interval] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5716
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5717
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5718
lemma inter_interval_mixed_eq_empty: fixes a :: "'a::ordered_euclidean_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5719
  assumes "{c<..<d} \<noteq> {}"  shows "{a<..<b} \<inter> {c .. d} = {} \<longleftrightarrow> {a<..<b} \<inter> {c<..<d} = {}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5720
  unfolding closure_open_interval[OF assms, THEN sym] unfolding open_inter_closure_eq_empty[OF open_interval] ..
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5721
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5722
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5723
(* Some stuff for half-infinite intervals too; FIXME: notation?  *)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5724
37673
f69f4b079275 generalize more lemmas from ordered_euclidean_space to euclidean_space
huffman
parents: 37649
diff changeset
  5725
lemma closed_interval_left: fixes b::"'a::euclidean_space"
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5726
  shows "closed {x::'a. \<forall>i\<in>Basis. x\<bullet>i \<le> b\<bullet>i}"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5727
proof-
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5728
  { fix i :: 'a assume i:"i\<in>Basis"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5729
    fix x::"'a" assume x:"\<forall>e>0. \<exists>x'\<in>{x. \<forall>i\<in>Basis. x \<bullet> i \<le> b \<bullet> i}. x' \<noteq> x \<and> dist x' x < e"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5730
    { assume "x\<bullet>i > b\<bullet>i"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5731
      then obtain y where "y \<bullet> i \<le> b \<bullet> i"  "y \<noteq> x"  "dist y x < x\<bullet>i - b\<bullet>i"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5732
        using x[THEN spec[where x="x\<bullet>i - b\<bullet>i"]] using i by auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5733
      hence False using Basis_le_norm[OF i, of "y - x"] unfolding dist_norm inner_simps using i 
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5734
        by auto }
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5735
    hence "x\<bullet>i \<le> b\<bullet>i" by(rule ccontr)auto  }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5736
  thus ?thesis unfolding closed_limpt unfolding islimpt_approachable by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5737
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5738
37673
f69f4b079275 generalize more lemmas from ordered_euclidean_space to euclidean_space
huffman
parents: 37649
diff changeset
  5739
lemma closed_interval_right: fixes a::"'a::euclidean_space"
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5740
  shows "closed {x::'a. \<forall>i\<in>Basis. a\<bullet>i \<le> x\<bullet>i}"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5741
proof-
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5742
  { fix i :: 'a assume i:"i\<in>Basis"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5743
    fix x::"'a" assume x:"\<forall>e>0. \<exists>x'\<in>{x. \<forall>i\<in>Basis. a \<bullet> i \<le> x \<bullet> i}. x' \<noteq> x \<and> dist x' x < e"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5744
    { assume "a\<bullet>i > x\<bullet>i"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5745
      then obtain y where "a \<bullet> i \<le> y \<bullet> i"  "y \<noteq> x"  "dist y x < a\<bullet>i - x\<bullet>i"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5746
        using x[THEN spec[where x="a\<bullet>i - x\<bullet>i"]] i by auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5747
      hence False using Basis_le_norm[OF i, of "y - x"] unfolding dist_norm inner_simps by auto }
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5748
    hence "a\<bullet>i \<le> x\<bullet>i" by(rule ccontr)auto  }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5749
  thus ?thesis unfolding closed_limpt unfolding islimpt_approachable by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5750
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5751
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5752
lemma open_box: "open (box a b)"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5753
proof -
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5754
  have "open (\<Inter>i\<in>Basis. (op \<bullet> i) -` {a \<bullet> i <..< b \<bullet> i})"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5755
    by (auto intro!: continuous_open_vimage continuous_inner continuous_at_id continuous_const)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5756
  also have "(\<Inter>i\<in>Basis. (op \<bullet> i) -` {a \<bullet> i <..< b \<bullet> i}) = box a b"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5757
    by (auto simp add: box_def inner_commute)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5758
  finally show ?thesis .
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5759
qed
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5760
50881
ae630bab13da renamed countable_basis_space to second_countable_topology
hoelzl
parents: 50526
diff changeset
  5761
instance euclidean_space \<subseteq> second_countable_topology
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  5762
proof
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5763
  def a \<equiv> "\<lambda>f :: 'a \<Rightarrow> (real \<times> real). \<Sum>i\<in>Basis. fst (f i) *\<^sub>R i"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5764
  then have a: "\<And>f. (\<Sum>i\<in>Basis. fst (f i) *\<^sub>R i) = a f" by simp
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5765
  def b \<equiv> "\<lambda>f :: 'a \<Rightarrow> (real \<times> real). \<Sum>i\<in>Basis. snd (f i) *\<^sub>R i"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5766
  then have b: "\<And>f. (\<Sum>i\<in>Basis. snd (f i) *\<^sub>R i) = b f" by simp
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5767
  def B \<equiv> "(\<lambda>f. box (a f) (b f)) ` (Basis \<rightarrow>\<^isub>E (\<rat> \<times> \<rat>))"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5768
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5769
  have "countable B" unfolding B_def 
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5770
    by (intro countable_image countable_PiE finite_Basis countable_SIGMA countable_rat)
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
  5771
  moreover
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5772
  have "Ball B open" by (simp add: B_def open_box)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5773
  moreover have "(\<forall>A. open A \<longrightarrow> (\<exists>B'\<subseteq>B. \<Union>B' = A))"
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  5774
  proof safe
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5775
    fix A::"'a set" assume "open A"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5776
    show "\<exists>B'\<subseteq>B. \<Union>B' = A"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5777
      apply (rule exI[of _ "{b\<in>B. b \<subseteq> A}"])
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5778
      apply (subst (3) open_UNION_box[OF `open A`])
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5779
      apply (auto simp add: a b B_def)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5780
      done
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  5781
  qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  5782
  ultimately
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
  5783
  show "\<exists>B::'a set set. countable B \<and> topological_basis B" unfolding topological_basis_def by blast
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  5784
qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  5785
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  5786
instance ordered_euclidean_space \<subseteq> polish_space ..
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  5787
36439
a65320184de9 move intervals section heading
huffman
parents: 36438
diff changeset
  5788
text {* Intervals in general, including infinite and mixtures of open and closed. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5789
37732
6432bf0d7191 generalize type of is_interval to class euclidean_space
huffman
parents: 37680
diff changeset
  5790
definition "is_interval (s::('a::euclidean_space) set) \<longleftrightarrow>
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5791
  (\<forall>a\<in>s. \<forall>b\<in>s. \<forall>x. (\<forall>i\<in>Basis. ((a\<bullet>i \<le> x\<bullet>i \<and> x\<bullet>i \<le> b\<bullet>i) \<or> (b\<bullet>i \<le> x\<bullet>i \<and> x\<bullet>i \<le> a\<bullet>i))) \<longrightarrow> x \<in> s)"
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5792
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5793
lemma is_interval_interval: "is_interval {a .. b::'a::ordered_euclidean_space}" (is ?th1)
39086
c4b809e57fe0 preimages of open sets over continuous function are open
hoelzl
parents: 38656
diff changeset
  5794
  "is_interval {a<..<b}" (is ?th2) proof -
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5795
  show ?th1 ?th2  unfolding is_interval_def mem_interval Ball_def atLeastAtMost_iff
44584
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  5796
    by(meson order_trans le_less_trans less_le_trans less_trans)+ qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5797
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5798
lemma is_interval_empty:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5799
 "is_interval {}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5800
  unfolding is_interval_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5801
  by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5802
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5803
lemma is_interval_univ:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5804
 "is_interval UNIV"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5805
  unfolding is_interval_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5806
  by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5807
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  5808
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  5809
subsection {* Closure of halfspaces and hyperplanes *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5810
44219
f738e3200e24 generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
huffman
parents: 44216
diff changeset
  5811
lemma isCont_open_vimage:
f738e3200e24 generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
huffman
parents: 44216
diff changeset
  5812
  assumes "\<And>x. isCont f x" and "open s" shows "open (f -` s)"
f738e3200e24 generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
huffman
parents: 44216
diff changeset
  5813
proof -
f738e3200e24 generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
huffman
parents: 44216
diff changeset
  5814
  from assms(1) have "continuous_on UNIV f"
f738e3200e24 generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
huffman
parents: 44216
diff changeset
  5815
    unfolding isCont_def continuous_on_def within_UNIV by simp
f738e3200e24 generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
huffman
parents: 44216
diff changeset
  5816
  hence "open {x \<in> UNIV. f x \<in> s}"
f738e3200e24 generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
huffman
parents: 44216
diff changeset
  5817
    using open_UNIV `open s` by (rule continuous_open_preimage)
f738e3200e24 generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
huffman
parents: 44216
diff changeset
  5818
  thus "open (f -` s)"
f738e3200e24 generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
huffman
parents: 44216
diff changeset
  5819
    by (simp add: vimage_def)
f738e3200e24 generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
huffman
parents: 44216
diff changeset
  5820
qed
f738e3200e24 generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
huffman
parents: 44216
diff changeset
  5821
f738e3200e24 generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
huffman
parents: 44216
diff changeset
  5822
lemma isCont_closed_vimage:
f738e3200e24 generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
huffman
parents: 44216
diff changeset
  5823
  assumes "\<And>x. isCont f x" and "closed s" shows "closed (f -` s)"
f738e3200e24 generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
huffman
parents: 44216
diff changeset
  5824
  using assms unfolding closed_def vimage_Compl [symmetric]
f738e3200e24 generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
huffman
parents: 44216
diff changeset
  5825
  by (rule isCont_open_vimage)
f738e3200e24 generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
huffman
parents: 44216
diff changeset
  5826
44213
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5827
lemma open_Collect_less:
44219
f738e3200e24 generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
huffman
parents: 44216
diff changeset
  5828
  fixes f g :: "'a::topological_space \<Rightarrow> real"
f738e3200e24 generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
huffman
parents: 44216
diff changeset
  5829
  assumes f: "\<And>x. isCont f x"
f738e3200e24 generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
huffman
parents: 44216
diff changeset
  5830
  assumes g: "\<And>x. isCont g x"
44213
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5831
  shows "open {x. f x < g x}"
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5832
proof -
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5833
  have "open ((\<lambda>x. g x - f x) -` {0<..})"
44219
f738e3200e24 generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
huffman
parents: 44216
diff changeset
  5834
    using isCont_diff [OF g f] open_real_greaterThan
f738e3200e24 generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
huffman
parents: 44216
diff changeset
  5835
    by (rule isCont_open_vimage)
44213
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5836
  also have "((\<lambda>x. g x - f x) -` {0<..}) = {x. f x < g x}"
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5837
    by auto
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5838
  finally show ?thesis .
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5839
qed
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5840
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5841
lemma closed_Collect_le:
44219
f738e3200e24 generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
huffman
parents: 44216
diff changeset
  5842
  fixes f g :: "'a::topological_space \<Rightarrow> real"
f738e3200e24 generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
huffman
parents: 44216
diff changeset
  5843
  assumes f: "\<And>x. isCont f x"
f738e3200e24 generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
huffman
parents: 44216
diff changeset
  5844
  assumes g: "\<And>x. isCont g x"
44213
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5845
  shows "closed {x. f x \<le> g x}"
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5846
proof -
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5847
  have "closed ((\<lambda>x. g x - f x) -` {0..})"
44219
f738e3200e24 generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
huffman
parents: 44216
diff changeset
  5848
    using isCont_diff [OF g f] closed_real_atLeast
f738e3200e24 generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
huffman
parents: 44216
diff changeset
  5849
    by (rule isCont_closed_vimage)
44213
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5850
  also have "((\<lambda>x. g x - f x) -` {0..}) = {x. f x \<le> g x}"
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5851
    by auto
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5852
  finally show ?thesis .
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5853
qed
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5854
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5855
lemma closed_Collect_eq:
44219
f738e3200e24 generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
huffman
parents: 44216
diff changeset
  5856
  fixes f g :: "'a::topological_space \<Rightarrow> 'b::t2_space"
f738e3200e24 generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
huffman
parents: 44216
diff changeset
  5857
  assumes f: "\<And>x. isCont f x"
f738e3200e24 generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
huffman
parents: 44216
diff changeset
  5858
  assumes g: "\<And>x. isCont g x"
44213
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5859
  shows "closed {x. f x = g x}"
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5860
proof -
44216
903bfe95fece generalized lemma closed_Collect_eq
huffman
parents: 44213
diff changeset
  5861
  have "open {(x::'b, y::'b). x \<noteq> y}"
903bfe95fece generalized lemma closed_Collect_eq
huffman
parents: 44213
diff changeset
  5862
    unfolding open_prod_def by (auto dest!: hausdorff)
903bfe95fece generalized lemma closed_Collect_eq
huffman
parents: 44213
diff changeset
  5863
  hence "closed {(x::'b, y::'b). x = y}"
903bfe95fece generalized lemma closed_Collect_eq
huffman
parents: 44213
diff changeset
  5864
    unfolding closed_def split_def Collect_neg_eq .
44219
f738e3200e24 generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
huffman
parents: 44216
diff changeset
  5865
  with isCont_Pair [OF f g]
44216
903bfe95fece generalized lemma closed_Collect_eq
huffman
parents: 44213
diff changeset
  5866
  have "closed ((\<lambda>x. (f x, g x)) -` {(x, y). x = y})"
44219
f738e3200e24 generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
huffman
parents: 44216
diff changeset
  5867
    by (rule isCont_closed_vimage)
44216
903bfe95fece generalized lemma closed_Collect_eq
huffman
parents: 44213
diff changeset
  5868
  also have "\<dots> = {x. f x = g x}" by auto
44213
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5869
  finally show ?thesis .
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5870
qed
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5871
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5872
lemma continuous_at_inner: "continuous (at x) (inner a)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5873
  unfolding continuous_at by (intro tendsto_intros)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5874
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5875
lemma closed_halfspace_le: "closed {x. inner a x \<le> b}"
44233
aa74ce315bae add simp rules for isCont
huffman
parents: 44219
diff changeset
  5876
  by (simp add: closed_Collect_le)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5877
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5878
lemma closed_halfspace_ge: "closed {x. inner a x \<ge> b}"
44233
aa74ce315bae add simp rules for isCont
huffman
parents: 44219
diff changeset
  5879
  by (simp add: closed_Collect_le)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5880
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5881
lemma closed_hyperplane: "closed {x. inner a x = b}"
44233
aa74ce315bae add simp rules for isCont
huffman
parents: 44219
diff changeset
  5882
  by (simp add: closed_Collect_eq)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5883
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5884
lemma closed_halfspace_component_le:
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5885
  shows "closed {x::'a::euclidean_space. x\<bullet>i \<le> a}"
44233
aa74ce315bae add simp rules for isCont
huffman
parents: 44219
diff changeset
  5886
  by (simp add: closed_Collect_le)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5887
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5888
lemma closed_halfspace_component_ge:
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5889
  shows "closed {x::'a::euclidean_space. x\<bullet>i \<ge> a}"
44233
aa74ce315bae add simp rules for isCont
huffman
parents: 44219
diff changeset
  5890
  by (simp add: closed_Collect_le)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5891
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  5892
text {* Openness of halfspaces. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5893
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5894
lemma open_halfspace_lt: "open {x. inner a x < b}"
44233
aa74ce315bae add simp rules for isCont
huffman
parents: 44219
diff changeset
  5895
  by (simp add: open_Collect_less)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5896
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5897
lemma open_halfspace_gt: "open {x. inner a x > b}"
44233
aa74ce315bae add simp rules for isCont
huffman
parents: 44219
diff changeset
  5898
  by (simp add: open_Collect_less)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5899
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5900
lemma open_halfspace_component_lt:
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5901
  shows "open {x::'a::euclidean_space. x\<bullet>i < a}"
44233
aa74ce315bae add simp rules for isCont
huffman
parents: 44219
diff changeset
  5902
  by (simp add: open_Collect_less)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5903
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5904
lemma open_halfspace_component_gt:
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5905
  shows "open {x::'a::euclidean_space. x\<bullet>i > a}"
44233
aa74ce315bae add simp rules for isCont
huffman
parents: 44219
diff changeset
  5906
  by (simp add: open_Collect_less)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5907
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5908
text{* Instantiation for intervals on @{text ordered_euclidean_space} *}
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5909
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5910
lemma eucl_lessThan_eq_halfspaces:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5911
  fixes a :: "'a\<Colon>ordered_euclidean_space"
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5912
  shows "{..<a} = (\<Inter>i\<in>Basis. {x. x \<bullet> i < a \<bullet> i})"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5913
 by (auto simp: eucl_less[where 'a='a])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5914
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5915
lemma eucl_greaterThan_eq_halfspaces:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5916
  fixes a :: "'a\<Colon>ordered_euclidean_space"
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5917
  shows "{a<..} = (\<Inter>i\<in>Basis. {x. a \<bullet> i < x \<bullet> i})"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5918
 by (auto simp: eucl_less[where 'a='a])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5919
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5920
lemma eucl_atMost_eq_halfspaces:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5921
  fixes a :: "'a\<Colon>ordered_euclidean_space"
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5922
  shows "{.. a} = (\<Inter>i\<in>Basis. {x. x \<bullet> i \<le> a \<bullet> i})"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5923
 by (auto simp: eucl_le[where 'a='a])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5924
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5925
lemma eucl_atLeast_eq_halfspaces:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5926
  fixes a :: "'a\<Colon>ordered_euclidean_space"
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5927
  shows "{a ..} = (\<Inter>i\<in>Basis. {x. a \<bullet> i \<le> x \<bullet> i})"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5928
 by (auto simp: eucl_le[where 'a='a])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5929
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5930
lemma open_eucl_lessThan[simp, intro]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5931
  fixes a :: "'a\<Colon>ordered_euclidean_space"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5932
  shows "open {..< a}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5933
  by (auto simp: eucl_lessThan_eq_halfspaces open_halfspace_component_lt)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5934
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5935
lemma open_eucl_greaterThan[simp, intro]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5936
  fixes a :: "'a\<Colon>ordered_euclidean_space"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5937
  shows "open {a <..}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5938
  by (auto simp: eucl_greaterThan_eq_halfspaces open_halfspace_component_gt)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5939
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5940
lemma closed_eucl_atMost[simp, intro]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5941
  fixes a :: "'a\<Colon>ordered_euclidean_space"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5942
  shows "closed {.. a}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5943
  unfolding eucl_atMost_eq_halfspaces
44233
aa74ce315bae add simp rules for isCont
huffman
parents: 44219
diff changeset
  5944
  by (simp add: closed_INT closed_Collect_le)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5945
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5946
lemma closed_eucl_atLeast[simp, intro]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5947
  fixes a :: "'a\<Colon>ordered_euclidean_space"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5948
  shows "closed {a ..}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5949
  unfolding eucl_atLeast_eq_halfspaces
44233
aa74ce315bae add simp rules for isCont
huffman
parents: 44219
diff changeset
  5950
  by (simp add: closed_INT closed_Collect_le)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5951
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  5952
text {* This gives a simple derivation of limit component bounds. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5953
37673
f69f4b079275 generalize more lemmas from ordered_euclidean_space to euclidean_space
huffman
parents: 37649
diff changeset
  5954
lemma Lim_component_le: fixes f :: "'a \<Rightarrow> 'b::euclidean_space"
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5955
  assumes "(f ---> l) net" "\<not> (trivial_limit net)"  "eventually (\<lambda>x. f(x)\<bullet>i \<le> b) net"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5956
  shows "l\<bullet>i \<le> b"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5957
  by (rule tendsto_le[OF assms(2) tendsto_const tendsto_inner[OF assms(1) tendsto_const] assms(3)])
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5958
37673
f69f4b079275 generalize more lemmas from ordered_euclidean_space to euclidean_space
huffman
parents: 37649
diff changeset
  5959
lemma Lim_component_ge: fixes f :: "'a \<Rightarrow> 'b::euclidean_space"
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5960
  assumes "(f ---> l) net"  "\<not> (trivial_limit net)"  "eventually (\<lambda>x. b \<le> (f x)\<bullet>i) net"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5961
  shows "b \<le> l\<bullet>i"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5962
  by (rule tendsto_le[OF assms(2) tendsto_inner[OF assms(1) tendsto_const] tendsto_const assms(3)])
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5963
37673
f69f4b079275 generalize more lemmas from ordered_euclidean_space to euclidean_space
huffman
parents: 37649
diff changeset
  5964
lemma Lim_component_eq: fixes f :: "'a \<Rightarrow> 'b::euclidean_space"
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5965
  assumes net:"(f ---> l) net" "~(trivial_limit net)" and ev:"eventually (\<lambda>x. f(x)\<bullet>i = b) net"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5966
  shows "l\<bullet>i = b"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5967
  using ev[unfolded order_eq_iff eventually_conj_iff]
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5968
  using Lim_component_ge[OF net, of b i] and Lim_component_le[OF net, of i b] by auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5969
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5970
text{* Limits relative to a union.                                               *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5971
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5972
lemma eventually_within_Un:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5973
  "eventually P (net within (s \<union> t)) \<longleftrightarrow>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5974
    eventually P (net within s) \<and> eventually P (net within t)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5975
  unfolding Limits.eventually_within
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5976
  by (auto elim!: eventually_rev_mp)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5977
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5978
lemma Lim_within_union:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5979
 "(f ---> l) (net within (s \<union> t)) \<longleftrightarrow>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5980
  (f ---> l) (net within s) \<and> (f ---> l) (net within t)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5981
  unfolding tendsto_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5982
  by (auto simp add: eventually_within_Un)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5983
36442
b96d9dc6acca generalize more continuity lemmas
huffman
parents: 36441
diff changeset
  5984
lemma Lim_topological:
b96d9dc6acca generalize more continuity lemmas
huffman
parents: 36441
diff changeset
  5985
 "(f ---> l) net \<longleftrightarrow>
b96d9dc6acca generalize more continuity lemmas
huffman
parents: 36441
diff changeset
  5986
        trivial_limit net \<or>
b96d9dc6acca generalize more continuity lemmas
huffman
parents: 36441
diff changeset
  5987
        (\<forall>S. open S \<longrightarrow> l \<in> S \<longrightarrow> eventually (\<lambda>x. f x \<in> S) net)"
b96d9dc6acca generalize more continuity lemmas
huffman
parents: 36441
diff changeset
  5988
  unfolding tendsto_def trivial_limit_eq by auto
b96d9dc6acca generalize more continuity lemmas
huffman
parents: 36441
diff changeset
  5989
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5990
lemma continuous_on_union:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5991
  assumes "closed s" "closed t" "continuous_on s f" "continuous_on t f"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5992
  shows "continuous_on (s \<union> t) f"
36442
b96d9dc6acca generalize more continuity lemmas
huffman
parents: 36441
diff changeset
  5993
  using assms unfolding continuous_on Lim_within_union
b96d9dc6acca generalize more continuity lemmas
huffman
parents: 36441
diff changeset
  5994
  unfolding Lim_topological trivial_limit_within closed_limpt by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5995
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5996
lemma continuous_on_cases:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5997
  assumes "closed s" "closed t" "continuous_on s f" "continuous_on t g"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5998
          "\<forall>x. (x\<in>s \<and> \<not> P x) \<or> (x \<in> t \<and> P x) \<longrightarrow> f x = g x"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5999
  shows "continuous_on (s \<union> t) (\<lambda>x. if P x then f x else g x)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6000
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6001
  let ?h = "(\<lambda>x. if P x then f x else g x)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6002
  have "\<forall>x\<in>s. f x = (if P x then f x else g x)" using assms(5) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6003
  hence "continuous_on s ?h" using continuous_on_eq[of s f ?h] using assms(3) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6004
  moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6005
  have "\<forall>x\<in>t. g x = (if P x then f x else g x)" using assms(5) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6006
  hence "continuous_on t ?h" using continuous_on_eq[of t g ?h] using assms(4) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6007
  ultimately show ?thesis using continuous_on_union[OF assms(1,2), of ?h] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6008
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6009
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6010
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6011
text{* Some more convenient intermediate-value theorem formulations.             *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6012
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6013
lemma connected_ivt_hyperplane:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6014
  assumes "connected s" "x \<in> s" "y \<in> s" "inner a x \<le> b" "b \<le> inner a y"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6015
  shows "\<exists>z \<in> s. inner a z = b"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6016
proof(rule ccontr)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6017
  assume as:"\<not> (\<exists>z\<in>s. inner a z = b)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6018
  let ?A = "{x. inner a x < b}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6019
  let ?B = "{x. inner a x > b}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6020
  have "open ?A" "open ?B" using open_halfspace_lt and open_halfspace_gt by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6021
  moreover have "?A \<inter> ?B = {}" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6022
  moreover have "s \<subseteq> ?A \<union> ?B" using as by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6023
  ultimately show False using assms(1)[unfolded connected_def not_ex, THEN spec[where x="?A"], THEN spec[where x="?B"]] and assms(2-5) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6024
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6025
37673
f69f4b079275 generalize more lemmas from ordered_euclidean_space to euclidean_space
huffman
parents: 37649
diff changeset
  6026
lemma connected_ivt_component: fixes x::"'a::euclidean_space" shows
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  6027
 "connected s \<Longrightarrow> x \<in> s \<Longrightarrow> y \<in> s \<Longrightarrow> x\<bullet>k \<le> a \<Longrightarrow> a \<le> y\<bullet>k \<Longrightarrow> (\<exists>z\<in>s.  z\<bullet>k = a)"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  6028
  using connected_ivt_hyperplane[of s x y "k::'a" a] by (auto simp: inner_commute)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6029
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  6030
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  6031
subsection {* Homeomorphisms *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6032
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6033
definition "homeomorphism s t f g \<equiv>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6034
     (\<forall>x\<in>s. (g(f x) = x)) \<and> (f ` s = t) \<and> continuous_on s f \<and>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6035
     (\<forall>y\<in>t. (f(g y) = y)) \<and> (g ` t = s) \<and> continuous_on t g"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6036
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6037
definition
50898
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  6038
  homeomorphic :: "'a::topological_space set \<Rightarrow> 'b::topological_space set \<Rightarrow> bool"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6039
    (infixr "homeomorphic" 60) where
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6040
  homeomorphic_def: "s homeomorphic t \<equiv> (\<exists>f g. homeomorphism s t f g)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6041
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6042
lemma homeomorphic_refl: "s homeomorphic s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6043
  unfolding homeomorphic_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6044
  unfolding homeomorphism_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6045
  using continuous_on_id
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6046
  apply(rule_tac x = "(\<lambda>x. x)" in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6047
  apply(rule_tac x = "(\<lambda>x. x)" in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6048
  by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6049
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6050
lemma homeomorphic_sym:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6051
 "s homeomorphic t \<longleftrightarrow> t homeomorphic s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6052
unfolding homeomorphic_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6053
unfolding homeomorphism_def
33324
51eb2ffa2189 Tidied up some very ugly proofs
paulson
parents: 33270
diff changeset
  6054
by blast 
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6055
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6056
lemma homeomorphic_trans:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6057
  assumes "s homeomorphic t" "t homeomorphic u" shows "s homeomorphic u"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6058
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6059
  obtain f1 g1 where fg1:"\<forall>x\<in>s. g1 (f1 x) = x"  "f1 ` s = t" "continuous_on s f1" "\<forall>y\<in>t. f1 (g1 y) = y" "g1 ` t = s" "continuous_on t g1"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6060
    using assms(1) unfolding homeomorphic_def homeomorphism_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6061
  obtain f2 g2 where fg2:"\<forall>x\<in>t. g2 (f2 x) = x"  "f2 ` t = u" "continuous_on t f2" "\<forall>y\<in>u. f2 (g2 y) = y" "g2 ` u = t" "continuous_on u g2"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6062
    using assms(2) unfolding homeomorphic_def homeomorphism_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6063
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6064
  { fix x assume "x\<in>s" hence "(g1 \<circ> g2) ((f2 \<circ> f1) x) = x" using fg1(1)[THEN bspec[where x=x]] and fg2(1)[THEN bspec[where x="f1 x"]] and fg1(2) by auto }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6065
  moreover have "(f2 \<circ> f1) ` s = u" using fg1(2) fg2(2) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6066
  moreover have "continuous_on s (f2 \<circ> f1)" using continuous_on_compose[OF fg1(3)] and fg2(3) unfolding fg1(2) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6067
  moreover { fix y assume "y\<in>u" hence "(f2 \<circ> f1) ((g1 \<circ> g2) y) = y" using fg2(4)[THEN bspec[where x=y]] and fg1(4)[THEN bspec[where x="g2 y"]] and fg2(5) by auto }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6068
  moreover have "(g1 \<circ> g2) ` u = s" using fg1(5) fg2(5) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6069
  moreover have "continuous_on u (g1 \<circ> g2)" using continuous_on_compose[OF fg2(6)] and fg1(6)  unfolding fg2(5) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6070
  ultimately show ?thesis unfolding homeomorphic_def homeomorphism_def apply(rule_tac x="f2 \<circ> f1" in exI) apply(rule_tac x="g1 \<circ> g2" in exI) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6071
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6072
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6073
lemma homeomorphic_minimal:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6074
 "s homeomorphic t \<longleftrightarrow>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6075
    (\<exists>f g. (\<forall>x\<in>s. f(x) \<in> t \<and> (g(f(x)) = x)) \<and>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6076
           (\<forall>y\<in>t. g(y) \<in> s \<and> (f(g(y)) = y)) \<and>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6077
           continuous_on s f \<and> continuous_on t g)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6078
unfolding homeomorphic_def homeomorphism_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6079
apply auto apply (rule_tac x=f in exI) apply (rule_tac x=g in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6080
apply auto apply (rule_tac x=f in exI) apply (rule_tac x=g in exI) apply auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6081
unfolding image_iff
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6082
apply(erule_tac x="g x" in ballE) apply(erule_tac x="x" in ballE)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6083
apply auto apply(rule_tac x="g x" in bexI) apply auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6084
apply(erule_tac x="f x" in ballE) apply(erule_tac x="x" in ballE)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6085
apply auto apply(rule_tac x="f x" in bexI) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6086
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  6087
text {* Relatively weak hypotheses if a set is compact. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6088
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6089
lemma homeomorphism_compact:
50898
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  6090
  fixes f :: "'a::topological_space \<Rightarrow> 'b::t2_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6091
  assumes "compact s" "continuous_on s f"  "f ` s = t"  "inj_on f s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6092
  shows "\<exists>g. homeomorphism s t f g"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6093
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6094
  def g \<equiv> "\<lambda>x. SOME y. y\<in>s \<and> f y = x"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6095
  have g:"\<forall>x\<in>s. g (f x) = x" using assms(3) assms(4)[unfolded inj_on_def] unfolding g_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6096
  { fix y assume "y\<in>t"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6097
    then obtain x where x:"f x = y" "x\<in>s" using assms(3) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6098
    hence "g (f x) = x" using g by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6099
    hence "f (g y) = y" unfolding x(1)[THEN sym] by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6100
  hence g':"\<forall>x\<in>t. f (g x) = x" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6101
  moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6102
  { fix x
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6103
    have "x\<in>s \<Longrightarrow> x \<in> g ` t" using g[THEN bspec[where x=x]] unfolding image_iff using assms(3) by(auto intro!: bexI[where x="f x"])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6104
    moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6105
    { assume "x\<in>g ` t"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6106
      then obtain y where y:"y\<in>t" "g y = x" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6107
      then obtain x' where x':"x'\<in>s" "f x' = y" using assms(3) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6108
      hence "x \<in> s" unfolding g_def using someI2[of "\<lambda>b. b\<in>s \<and> f b = y" x' "\<lambda>x. x\<in>s"] unfolding y(2)[THEN sym] and g_def by auto }
36362
06475a1547cb fix lots of looping simp calls and other warnings
huffman
parents: 36360
diff changeset
  6109
    ultimately have "x\<in>s \<longleftrightarrow> x \<in> g ` t" ..  }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6110
  hence "g ` t = s" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6111
  ultimately
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6112
  show ?thesis unfolding homeomorphism_def homeomorphic_def
44647
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  6113
    apply(rule_tac x=g in exI) using g and assms(3) and continuous_on_inv[OF assms(2,1), of g, unfolded assms(3)] and assms(2) by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6114
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6115
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6116
lemma homeomorphic_compact:
50898
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  6117
  fixes f :: "'a::topological_space \<Rightarrow> 'b::t2_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6118
  shows "compact s \<Longrightarrow> continuous_on s f \<Longrightarrow> (f ` s = t) \<Longrightarrow> inj_on f s
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6119
          \<Longrightarrow> s homeomorphic t"
37486
b993fac7985b beta-eta was too much, because it transformed SOME x. P x into Eps P, which caused problems later;
blanchet
parents: 37452
diff changeset
  6120
  unfolding homeomorphic_def by (metis homeomorphism_compact)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6121
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6122
text{* Preservation of topological properties.                                   *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6123
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6124
lemma homeomorphic_compactness:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6125
 "s homeomorphic t ==> (compact s \<longleftrightarrow> compact t)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6126
unfolding homeomorphic_def homeomorphism_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6127
by (metis compact_continuous_image)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6128
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6129
text{* Results on translation, scaling etc.                                      *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6130
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6131
lemma homeomorphic_scaling:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6132
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6133
  assumes "c \<noteq> 0"  shows "s homeomorphic ((\<lambda>x. c *\<^sub>R x) ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6134
  unfolding homeomorphic_minimal
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6135
  apply(rule_tac x="\<lambda>x. c *\<^sub>R x" in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6136
  apply(rule_tac x="\<lambda>x. (1 / c) *\<^sub>R x" in exI)
44531
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  6137
  using assms by (auto simp add: continuous_on_intros)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6138
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6139
lemma homeomorphic_translation:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6140
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6141
  shows "s homeomorphic ((\<lambda>x. a + x) ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6142
  unfolding homeomorphic_minimal
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6143
  apply(rule_tac x="\<lambda>x. a + x" in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6144
  apply(rule_tac x="\<lambda>x. -a + x" in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6145
  using continuous_on_add[OF continuous_on_const continuous_on_id] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6146
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6147
lemma homeomorphic_affinity:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6148
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6149
  assumes "c \<noteq> 0"  shows "s homeomorphic ((\<lambda>x. a + c *\<^sub>R x) ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6150
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6151
  have *:"op + a ` op *\<^sub>R c ` s = (\<lambda>x. a + c *\<^sub>R x) ` s" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6152
  show ?thesis
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6153
    using homeomorphic_trans
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6154
    using homeomorphic_scaling[OF assms, of s]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6155
    using homeomorphic_translation[of "(\<lambda>x. c *\<^sub>R x) ` s" a] unfolding * by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6156
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6157
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6158
lemma homeomorphic_balls:
50898
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  6159
  fixes a b ::"'a::real_normed_vector"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6160
  assumes "0 < d"  "0 < e"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6161
  shows "(ball a d) homeomorphic  (ball b e)" (is ?th)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6162
        "(cball a d) homeomorphic (cball b e)" (is ?cth)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6163
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6164
  show ?th unfolding homeomorphic_minimal
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6165
    apply(rule_tac x="\<lambda>x. b + (e/d) *\<^sub>R (x - a)" in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6166
    apply(rule_tac x="\<lambda>x. a + (d/e) *\<^sub>R (x - b)" in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6167
    using assms apply (auto simp add: dist_commute)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6168
    unfolding dist_norm
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6169
    apply (auto simp add: pos_divide_less_eq mult_strict_left_mono)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6170
    unfolding continuous_on
36659
f794e92784aa adapt to removed premise on tendsto lemma (cf. 88f0125c3bd2)
huffman
parents: 36623
diff changeset
  6171
    by (intro ballI tendsto_intros, simp)+
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6172
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6173
  show ?cth unfolding homeomorphic_minimal
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6174
    apply(rule_tac x="\<lambda>x. b + (e/d) *\<^sub>R (x - a)" in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6175
    apply(rule_tac x="\<lambda>x. a + (d/e) *\<^sub>R (x - b)" in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6176
    using assms apply (auto simp add: dist_commute)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6177
    unfolding dist_norm
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6178
    apply (auto simp add: pos_divide_le_eq)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6179
    unfolding continuous_on
36659
f794e92784aa adapt to removed premise on tendsto lemma (cf. 88f0125c3bd2)
huffman
parents: 36623
diff changeset
  6180
    by (intro ballI tendsto_intros, simp)+
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6181
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6182
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6183
text{* "Isometry" (up to constant bounds) of injective linear map etc.           *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6184
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6185
lemma cauchy_isometric:
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  6186
  fixes x :: "nat \<Rightarrow> 'a::euclidean_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6187
  assumes e:"0 < e" and s:"subspace s" and f:"bounded_linear f" and normf:"\<forall>x\<in>s. norm(f x) \<ge> e * norm(x)" and xs:"\<forall>n::nat. x n \<in> s" and cf:"Cauchy(f o x)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6188
  shows "Cauchy x"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6189
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6190
  interpret f: bounded_linear f by fact
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6191
  { fix d::real assume "d>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6192
    then obtain N where N:"\<forall>n\<ge>N. norm (f (x n) - f (x N)) < e * d"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6193
      using cf[unfolded cauchy o_def dist_norm, THEN spec[where x="e*d"]] and e and mult_pos_pos[of e d] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6194
    { fix n assume "n\<ge>N"
45270
d5b5c9259afd fix bug in cancel_factor simprocs so they will work on goals like 'x * y < x * z' where the common term is already on the left
huffman
parents: 45051
diff changeset
  6195
      have "e * norm (x n - x N) \<le> norm (f (x n - x N))"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6196
        using subspace_sub[OF s, of "x n" "x N"] using xs[THEN spec[where x=N]] and xs[THEN spec[where x=n]]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6197
        using normf[THEN bspec[where x="x n - x N"]] by auto
45270
d5b5c9259afd fix bug in cancel_factor simprocs so they will work on goals like 'x * y < x * z' where the common term is already on the left
huffman
parents: 45051
diff changeset
  6198
      also have "norm (f (x n - x N)) < e * d"
d5b5c9259afd fix bug in cancel_factor simprocs so they will work on goals like 'x * y < x * z' where the common term is already on the left
huffman
parents: 45051
diff changeset
  6199
        using `N \<le> n` N unfolding f.diff[THEN sym] by auto
d5b5c9259afd fix bug in cancel_factor simprocs so they will work on goals like 'x * y < x * z' where the common term is already on the left
huffman
parents: 45051
diff changeset
  6200
      finally have "norm (x n - x N) < d" using `e>0` by simp }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6201
    hence "\<exists>N. \<forall>n\<ge>N. norm (x n - x N) < d" by auto }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6202
  thus ?thesis unfolding cauchy and dist_norm by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6203
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6204
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6205
lemma complete_isometric_image:
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  6206
  fixes f :: "'a::euclidean_space => 'b::euclidean_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6207
  assumes "0 < e" and s:"subspace s" and f:"bounded_linear f" and normf:"\<forall>x\<in>s. norm(f x) \<ge> e * norm(x)" and cs:"complete s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6208
  shows "complete(f ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6209
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6210
  { fix g assume as:"\<forall>n::nat. g n \<in> f ` s" and cfg:"Cauchy g"
33324
51eb2ffa2189 Tidied up some very ugly proofs
paulson
parents: 33270
diff changeset
  6211
    then obtain x where "\<forall>n. x n \<in> s \<and> g n = f (x n)" 
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6212
      using choice[of "\<lambda> n xa. xa \<in> s \<and> g n = f xa"] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6213
    hence x:"\<forall>n. x n \<in> s"  "\<forall>n. g n = f (x n)" by auto
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
  6214
    hence "f \<circ> x = g" unfolding fun_eq_iff by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6215
    then obtain l where "l\<in>s" and l:"(x ---> l) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6216
      using cs[unfolded complete_def, THEN spec[where x="x"]]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6217
      using cauchy_isometric[OF `0<e` s f normf] and cfg and x(1) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6218
    hence "\<exists>l\<in>f ` s. (g ---> l) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6219
      using linear_continuous_at[OF f, unfolded continuous_at_sequentially, THEN spec[where x=x], of l]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6220
      unfolding `f \<circ> x = g` by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6221
  thus ?thesis unfolding complete_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6222
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6223
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  6224
lemma injective_imp_isometric: fixes f::"'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6225
  assumes s:"closed s"  "subspace s"  and f:"bounded_linear f" "\<forall>x\<in>s. (f x = 0) \<longrightarrow> (x = 0)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6226
  shows "\<exists>e>0. \<forall>x\<in>s. norm (f x) \<ge> e * norm(x)"
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  6227
proof(cases "s \<subseteq> {0::'a}")
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6228
  case True
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6229
  { fix x assume "x \<in> s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6230
    hence "x = 0" using True by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6231
    hence "norm x \<le> norm (f x)" by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6232
  thus ?thesis by(auto intro!: exI[where x=1])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6233
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6234
  interpret f: bounded_linear f by fact
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6235
  case False
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6236
  then obtain a where a:"a\<noteq>0" "a\<in>s" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6237
  from False have "s \<noteq> {}" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6238
  let ?S = "{f x| x. (x \<in> s \<and> norm x = norm a)}"
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  6239
  let ?S' = "{x::'a. x\<in>s \<and> norm x = norm a}"
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  6240
  let ?S'' = "{x::'a. norm x = norm a}"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6241
36362
06475a1547cb fix lots of looping simp calls and other warnings
huffman
parents: 36360
diff changeset
  6242
  have "?S'' = frontier(cball 0 (norm a))" unfolding frontier_cball and dist_norm by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6243
  hence "compact ?S''" using compact_frontier[OF compact_cball, of 0 "norm a"] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6244
  moreover have "?S' = s \<inter> ?S''" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6245
  ultimately have "compact ?S'" using closed_inter_compact[of s ?S''] using s(1) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6246
  moreover have *:"f ` ?S' = ?S" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6247
  ultimately have "compact ?S" using compact_continuous_image[OF linear_continuous_on[OF f(1)], of ?S'] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6248
  hence "closed ?S" using compact_imp_closed by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6249
  moreover have "?S \<noteq> {}" using a by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6250
  ultimately obtain b' where "b'\<in>?S" "\<forall>y\<in>?S. norm b' \<le> norm y" using distance_attains_inf[of ?S 0] unfolding dist_0_norm by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6251
  then obtain b where "b\<in>s" and ba:"norm b = norm a" and b:"\<forall>x\<in>{x \<in> s. norm x = norm a}. norm (f b) \<le> norm (f x)" unfolding *[THEN sym] unfolding image_iff by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6252
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6253
  let ?e = "norm (f b) / norm b"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6254
  have "norm b > 0" using ba and a and norm_ge_zero by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6255
  moreover have "norm (f b) > 0" using f(2)[THEN bspec[where x=b], OF `b\<in>s`] using `norm b >0` unfolding zero_less_norm_iff by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6256
  ultimately have "0 < norm (f b) / norm b" by(simp only: divide_pos_pos)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6257
  moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6258
  { fix x assume "x\<in>s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6259
    hence "norm (f b) / norm b * norm x \<le> norm (f x)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6260
    proof(cases "x=0")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6261
      case True thus "norm (f b) / norm b * norm x \<le> norm (f x)" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6262
    next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6263
      case False
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6264
      hence *:"0 < norm a / norm x" using `a\<noteq>0` unfolding zero_less_norm_iff[THEN sym] by(simp only: divide_pos_pos)
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  6265
      have "\<forall>c. \<forall>x\<in>s. c *\<^sub>R x \<in> s" using s[unfolded subspace_def] by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6266
      hence "(norm a / norm x) *\<^sub>R x \<in> {x \<in> s. norm x = norm a}" using `x\<in>s` and `x\<noteq>0` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6267
      thus "norm (f b) / norm b * norm x \<le> norm (f x)" using b[THEN bspec[where x="(norm a / norm x) *\<^sub>R x"]]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6268
        unfolding f.scaleR and ba using `x\<noteq>0` `a\<noteq>0`
36778
739a9379e29b avoid using real-specific versions of generic lemmas
huffman
parents: 36669
diff changeset
  6269
        by (auto simp add: mult_commute pos_le_divide_eq pos_divide_le_eq)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6270
    qed }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6271
  ultimately
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6272
  show ?thesis by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6273
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6274
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6275
lemma closed_injective_image_subspace:
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  6276
  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6277
  assumes "subspace s" "bounded_linear f" "\<forall>x\<in>s. f x = 0 --> x = 0" "closed s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6278
  shows "closed(f ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6279
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6280
  obtain e where "e>0" and e:"\<forall>x\<in>s. e * norm x \<le> norm (f x)" using injective_imp_isometric[OF assms(4,1,2,3)] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6281
  show ?thesis using complete_isometric_image[OF `e>0` assms(1,2) e] and assms(4)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6282
    unfolding complete_eq_closed[THEN sym] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6283
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6284
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  6285
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  6286
subsection {* Some properties of a canonical subspace *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6287
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6288
lemma subspace_substandard:
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  6289
  "subspace {x::'a::euclidean_space. (\<forall>i\<in>Basis. P i \<longrightarrow> x\<bullet>i = 0)}"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  6290
  unfolding subspace_def by (auto simp: inner_add_left)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6291
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6292
lemma closed_substandard:
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  6293
 "closed {x::'a::euclidean_space. \<forall>i\<in>Basis. P i --> x\<bullet>i = 0}" (is "closed ?A")
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6294
proof-
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  6295
  let ?D = "{i\<in>Basis. P i}"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  6296
  have "closed (\<Inter>i\<in>?D. {x::'a. x\<bullet>i = 0})"
44457
d366fa5551ef declare euclidean_simps [simp] at the point they are proved;
huffman
parents: 44365
diff changeset
  6297
    by (simp add: closed_INT closed_Collect_eq)
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  6298
  also have "(\<Inter>i\<in>?D. {x::'a. x\<bullet>i = 0}) = ?A"
44457
d366fa5551ef declare euclidean_simps [simp] at the point they are proved;
huffman
parents: 44365
diff changeset
  6299
    by auto
d366fa5551ef declare euclidean_simps [simp] at the point they are proved;
huffman
parents: 44365
diff changeset
  6300
  finally show "closed ?A" .
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6301
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6302
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  6303
lemma dim_substandard: assumes d: "d \<subseteq> Basis"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  6304
  shows "dim {x::'a::euclidean_space. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0} = card d" (is "dim ?A = _")
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6305
proof-
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  6306
  let ?D = "Basis :: 'a set"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  6307
  have "d \<subseteq> ?A" using d by (auto simp: inner_Basis)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6308
  moreover
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  6309
  { fix x::"'a" assume "x \<in> ?A"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  6310
    hence "finite d" "x \<in> ?A" using assms by(auto intro: finite_subset[OF _ finite_Basis])
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  6311
    from this d have "x \<in> span d"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6312
    proof(induct d arbitrary: x)
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  6313
      case empty hence "x=0" apply(rule_tac euclidean_eqI) by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6314
      thus ?case using subspace_0[OF subspace_span[of "{}"]] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6315
    next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6316
      case (insert k F)
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  6317
      hence *:"\<forall>i\<in>Basis. i \<notin> insert k F \<longrightarrow> x \<bullet> i = 0" by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6318
      have **:"F \<subseteq> insert k F" by auto
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  6319
      def y \<equiv> "x - (x\<bullet>k) *\<^sub>R k"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  6320
      have y:"x = y + (x\<bullet>k) *\<^sub>R k" unfolding y_def by auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  6321
      { fix i assume i': "i \<notin> F" "i \<in> Basis"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  6322
        hence "y \<bullet> i = 0" unfolding y_def 
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  6323
          using *[THEN bspec[where x=i]] insert by (auto simp: inner_simps inner_Basis) }
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  6324
      hence "y \<in> span F" using insert by auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  6325
      hence "y \<in> span (insert k F)"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  6326
        using span_mono[of F "insert k F"] using assms by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6327
      moreover
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  6328
      have "k \<in> span (insert k F)" by(rule span_superset, auto)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  6329
      hence "(x\<bullet>k) *\<^sub>R k \<in> span (insert k F)"
36593
fb69c8cd27bd define linear algebra concepts using scaleR instead of (op *s); generalized many lemmas, though a few theorems that used to work on type int^'n are a bit less general
huffman
parents: 36590
diff changeset
  6330
        using span_mul by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6331
      ultimately
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  6332
      have "y + (x\<bullet>k) *\<^sub>R k \<in> span (insert k F)"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6333
        using span_add by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6334
      thus ?case using y by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6335
    qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6336
  }
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  6337
  hence "?A \<subseteq> span d" by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6338
  moreover
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  6339
  { fix x assume "x \<in> d" hence "x \<in> ?D" using assms by auto  }
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  6340
  hence "independent d" using independent_mono[OF independent_Basis, of d] and assms by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6341
  moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6342
  have "d \<subseteq> ?D" unfolding subset_eq using assms by auto
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  6343
  ultimately show ?thesis using dim_unique[of d ?A] by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6344
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6345
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6346
text{* Hence closure and completeness of all subspaces.                          *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6347
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  6348
lemma ex_card: assumes "n \<le> card A" shows "\<exists>S\<subseteq>A. card S = n"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  6349
proof cases
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  6350
  assume "finite A"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  6351
  from ex_bij_betw_nat_finite[OF this] guess f ..
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  6352
  moreover with `n \<le> card A` have "{..< n} \<subseteq> {..< card A}" "inj_on f {..< n}"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  6353
    by (auto simp: bij_betw_def intro: subset_inj_on)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  6354
  ultimately have "f ` {..< n} \<subseteq> A" "card (f ` {..< n}) = n"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  6355
    by (auto simp: bij_betw_def card_image)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  6356
  then show ?thesis by blast
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  6357
next
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  6358
  assume "\<not> finite A" with `n \<le> card A` show ?thesis by force
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  6359
qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6360
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  6361
lemma closed_subspace: fixes s::"('a::euclidean_space) set"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6362
  assumes "subspace s" shows "closed s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6363
proof-
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  6364
  have "dim s \<le> card (Basis :: 'a set)" using dim_subset_UNIV by auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  6365
  with ex_card[OF this] obtain d :: "'a set" where t: "card d = dim s" and d: "d \<subseteq> Basis" by auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  6366
  let ?t = "{x::'a. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0}"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  6367
  have "\<exists>f. linear f \<and> f ` {x::'a. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0} = s \<and>
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  6368
      inj_on f {x::'a. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0}"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  6369
    using dim_substandard[of d] t d assms
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  6370
    by (intro subspace_isomorphism[OF subspace_substandard[of "\<lambda>i. i \<notin> d"]]) (auto simp: inner_Basis)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  6371
  then guess f by (elim exE conjE) note f = this
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  6372
  interpret f: bounded_linear f using f unfolding linear_conv_bounded_linear by auto
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  6373
  { fix x have "x\<in>?t \<Longrightarrow> f x = 0 \<Longrightarrow> x = 0" using f.zero d f(3)[THEN inj_onD, of x 0] by auto }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6374
  moreover have "closed ?t" using closed_substandard .
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6375
  moreover have "subspace ?t" using subspace_substandard .
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6376
  ultimately show ?thesis using closed_injective_image_subspace[of ?t f]
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  6377
    unfolding f(2) using f(1) unfolding linear_conv_bounded_linear by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6378
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6379
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6380
lemma complete_subspace:
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  6381
  fixes s :: "('a::euclidean_space) set" shows "subspace s ==> complete s"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6382
  using complete_eq_closed closed_subspace
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6383
  by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6384
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6385
lemma dim_closure:
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  6386
  fixes s :: "('a::euclidean_space) set"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6387
  shows "dim(closure s) = dim s" (is "?dc = ?d")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6388
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6389
  have "?dc \<le> ?d" using closure_minimal[OF span_inc, of s]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6390
    using closed_subspace[OF subspace_span, of s]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6391
    using dim_subset[of "closure s" "span s"] unfolding dim_span by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6392
  thus ?thesis using dim_subset[OF closure_subset, of s] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6393
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6394
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  6395
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  6396
subsection {* Affine transformations of intervals *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6397
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6398
lemma real_affinity_le:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34999
diff changeset
  6399
 "0 < (m::'a::linordered_field) ==> (m * x + c \<le> y \<longleftrightarrow> x \<le> inverse(m) * y + -(c / m))"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6400
  by (simp add: field_simps inverse_eq_divide)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6401
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6402
lemma real_le_affinity:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34999
diff changeset
  6403
 "0 < (m::'a::linordered_field) ==> (y \<le> m * x + c \<longleftrightarrow> inverse(m) * y + -(c / m) \<le> x)"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6404
  by (simp add: field_simps inverse_eq_divide)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6405
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6406
lemma real_affinity_lt:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34999
diff changeset
  6407
 "0 < (m::'a::linordered_field) ==> (m * x + c < y \<longleftrightarrow> x < inverse(m) * y + -(c / m))"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6408
  by (simp add: field_simps inverse_eq_divide)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6409
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6410
lemma real_lt_affinity:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34999
diff changeset
  6411
 "0 < (m::'a::linordered_field) ==> (y < m * x + c \<longleftrightarrow> inverse(m) * y + -(c / m) < x)"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6412
  by (simp add: field_simps inverse_eq_divide)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6413
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6414
lemma real_affinity_eq:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34999
diff changeset
  6415
 "(m::'a::linordered_field) \<noteq> 0 ==> (m * x + c = y \<longleftrightarrow> x = inverse(m) * y + -(c / m))"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6416
  by (simp add: field_simps inverse_eq_divide)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6417
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6418
lemma real_eq_affinity:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34999
diff changeset
  6419
 "(m::'a::linordered_field) \<noteq> 0 ==> (y = m * x + c  \<longleftrightarrow> inverse(m) * y + -(c / m) = x)"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6420
  by (simp add: field_simps inverse_eq_divide)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6421
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6422
lemma image_affinity_interval: fixes m::real
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  6423
  fixes a b c :: "'a::ordered_euclidean_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6424
  shows "(\<lambda>x. m *\<^sub>R x + c) ` {a .. b} =
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6425
            (if {a .. b} = {} then {}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6426
            else (if 0 \<le> m then {m *\<^sub>R a + c .. m *\<^sub>R b + c}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6427
            else {m *\<^sub>R b + c .. m *\<^sub>R a + c}))"
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  6428
proof(cases "m=0")  
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6429
  { fix x assume "x \<le> c" "c \<le> x"
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  6430
    hence "x=c" unfolding eucl_le[where 'a='a] apply-
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  6431
      apply(subst euclidean_eq_iff) by (auto intro: order_antisym) }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6432
  moreover case True
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  6433
  moreover have "c \<in> {m *\<^sub>R a + c..m *\<^sub>R b + c}" unfolding True by(auto simp add: eucl_le[where 'a='a])
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6434
  ultimately show ?thesis by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6435
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6436
  case False
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6437
  { fix y assume "a \<le> y" "y \<le> b" "m > 0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6438
    hence "m *\<^sub>R a + c \<le> m *\<^sub>R y + c"  "m *\<^sub>R y + c \<le> m *\<^sub>R b + c"
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  6439
      unfolding eucl_le[where 'a='a] by (auto simp: inner_simps)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6440
  } moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6441
  { fix y assume "a \<le> y" "y \<le> b" "m < 0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6442
    hence "m *\<^sub>R b + c \<le> m *\<^sub>R y + c"  "m *\<^sub>R y + c \<le> m *\<^sub>R a + c"
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  6443
      unfolding eucl_le[where 'a='a] by(auto simp add: mult_left_mono_neg inner_simps)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6444
  } moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6445
  { fix y assume "m > 0"  "m *\<^sub>R a + c \<le> y"  "y \<le> m *\<^sub>R b + c"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6446
    hence "y \<in> (\<lambda>x. m *\<^sub>R x + c) ` {a..b}"
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  6447
      unfolding image_iff Bex_def mem_interval eucl_le[where 'a='a]
44516
d9a496ae5d9d move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents: 44457
diff changeset
  6448
      apply (intro exI[where x="(1 / m) *\<^sub>R (y - c)"])
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  6449
      by(auto simp add: pos_le_divide_eq pos_divide_le_eq mult_commute diff_le_iff inner_simps)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6450
  } moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6451
  { fix y assume "m *\<^sub>R b + c \<le> y" "y \<le> m *\<^sub>R a + c" "m < 0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6452
    hence "y \<in> (\<lambda>x. m *\<^sub>R x + c) ` {a..b}"
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  6453
      unfolding image_iff Bex_def mem_interval eucl_le[where 'a='a]
44516
d9a496ae5d9d move everything related to 'norm' method into new theory file Norm_Arith.thy
huffman
parents: 44457
diff changeset
  6454
      apply (intro exI[where x="(1 / m) *\<^sub>R (y - c)"])
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  6455
      by(auto simp add: neg_le_divide_eq neg_divide_le_eq mult_commute diff_le_iff inner_simps)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6456
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6457
  ultimately show ?thesis using False by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6458
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6459
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  6460
lemma image_smult_interval:"(\<lambda>x. m *\<^sub>R (x::_::ordered_euclidean_space)) ` {a..b} =
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6461
  (if {a..b} = {} then {} else if 0 \<le> m then {m *\<^sub>R a..m *\<^sub>R b} else {m *\<^sub>R b..m *\<^sub>R a})"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6462
  using image_affinity_interval[of m 0 a b] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6463
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  6464
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  6465
subsection {* Banach fixed point theorem (not really topological...) *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6466
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6467
lemma banach_fix:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6468
  assumes s:"complete s" "s \<noteq> {}" and c:"0 \<le> c" "c < 1" and f:"(f ` s) \<subseteq> s" and
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6469
          lipschitz:"\<forall>x\<in>s. \<forall>y\<in>s. dist (f x) (f y) \<le> c * dist x y"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6470
  shows "\<exists>! x\<in>s. (f x = x)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6471
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6472
  have "1 - c > 0" using c by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6473
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6474
  from s(2) obtain z0 where "z0 \<in> s" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6475
  def z \<equiv> "\<lambda>n. (f ^^ n) z0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6476
  { fix n::nat
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6477
    have "z n \<in> s" unfolding z_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6478
    proof(induct n) case 0 thus ?case using `z0 \<in>s` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6479
    next case Suc thus ?case using f by auto qed }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6480
  note z_in_s = this
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6481
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6482
  def d \<equiv> "dist (z 0) (z 1)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6483
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6484
  have fzn:"\<And>n. f (z n) = z (Suc n)" unfolding z_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6485
  { fix n::nat
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6486
    have "dist (z n) (z (Suc n)) \<le> (c ^ n) * d"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6487
    proof(induct n)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6488
      case 0 thus ?case unfolding d_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6489
    next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6490
      case (Suc m)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6491
      hence "c * dist (z m) (z (Suc m)) \<le> c ^ Suc m * d"
38642
8fa437809c67 dropped type classes mult_mono and mult_mono1; tuned names of technical rule duplicates
haftmann
parents: 37732
diff changeset
  6492
        using `0 \<le> c` using mult_left_mono[of "dist (z m) (z (Suc m))" "c ^ m * d" c] by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6493
      thus ?case using lipschitz[THEN bspec[where x="z m"], OF z_in_s, THEN bspec[where x="z (Suc m)"], OF z_in_s]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6494
        unfolding fzn and mult_le_cancel_left by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6495
    qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6496
  } note cf_z = this
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6497
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6498
  { fix n m::nat
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6499
    have "(1 - c) * dist (z m) (z (m+n)) \<le> (c ^ m) * d * (1 - c ^ n)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6500
    proof(induct n)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6501
      case 0 show ?case by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6502
    next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6503
      case (Suc k)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6504
      have "(1 - c) * dist (z m) (z (m + Suc k)) \<le> (1 - c) * (dist (z m) (z (m + k)) + dist (z (m + k)) (z (Suc (m + k))))"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6505
        using dist_triangle and c by(auto simp add: dist_triangle)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6506
      also have "\<dots> \<le> (1 - c) * (dist (z m) (z (m + k)) + c ^ (m + k) * d)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6507
        using cf_z[of "m + k"] and c by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6508
      also have "\<dots> \<le> c ^ m * d * (1 - c ^ k) + (1 - c) * c ^ (m + k) * d"
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36336
diff changeset
  6509
        using Suc by (auto simp add: field_simps)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6510
      also have "\<dots> = (c ^ m) * (d * (1 - c ^ k) + (1 - c) * c ^ k * d)"
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36336
diff changeset
  6511
        unfolding power_add by (auto simp add: field_simps)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6512
      also have "\<dots> \<le> (c ^ m) * d * (1 - c ^ Suc k)"
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36336
diff changeset
  6513
        using c by (auto simp add: field_simps)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6514
      finally show ?case by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6515
    qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6516
  } note cf_z2 = this
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6517
  { fix e::real assume "e>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6518
    hence "\<exists>N. \<forall>m n. N \<le> m \<and> N \<le> n \<longrightarrow> dist (z m) (z n) < e"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6519
    proof(cases "d = 0")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6520
      case True
41863
e5104b436ea1 removed dependency on Dense_Linear_Order
boehmes
parents: 41413
diff changeset
  6521
      have *: "\<And>x. ((1 - c) * x \<le> 0) = (x \<le> 0)" using `1 - c > 0`
45051
c478d1876371 discontinued legacy theorem names from RealDef.thy
huffman
parents: 45031
diff changeset
  6522
        by (metis mult_zero_left mult_commute real_mult_le_cancel_iff1)
41863
e5104b436ea1 removed dependency on Dense_Linear_Order
boehmes
parents: 41413
diff changeset
  6523
      from True have "\<And>n. z n = z0" using cf_z2[of 0] and c unfolding z_def
e5104b436ea1 removed dependency on Dense_Linear_Order
boehmes
parents: 41413
diff changeset
  6524
        by (simp add: *)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6525
      thus ?thesis using `e>0` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6526
    next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6527
      case False hence "d>0" unfolding d_def using zero_le_dist[of "z 0" "z 1"]
36778
739a9379e29b avoid using real-specific versions of generic lemmas
huffman
parents: 36669
diff changeset
  6528
        by (metis False d_def less_le)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6529
      hence "0 < e * (1 - c) / d" using `e>0` and `1-c>0`
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6530
        using divide_pos_pos[of "e * (1 - c)" d] and mult_pos_pos[of e "1 - c"] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6531
      then obtain N where N:"c ^ N < e * (1 - c) / d" using real_arch_pow_inv[of "e * (1 - c) / d" c] and c by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6532
      { fix m n::nat assume "m>n" and as:"m\<ge>N" "n\<ge>N"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6533
        have *:"c ^ n \<le> c ^ N" using `n\<ge>N` and c using power_decreasing[OF `n\<ge>N`, of c] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6534
        have "1 - c ^ (m - n) > 0" using c and power_strict_mono[of c 1 "m - n"] using `m>n` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6535
        hence **:"d * (1 - c ^ (m - n)) / (1 - c) > 0"
36778
739a9379e29b avoid using real-specific versions of generic lemmas
huffman
parents: 36669
diff changeset
  6536
          using mult_pos_pos[OF `d>0`, of "1 - c ^ (m - n)"]
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6537
          using divide_pos_pos[of "d * (1 - c ^ (m - n))" "1 - c"]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6538
          using `0 < 1 - c` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6539
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6540
        have "dist (z m) (z n) \<le> c ^ n * d * (1 - c ^ (m - n)) / (1 - c)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6541
          using cf_z2[of n "m - n"] and `m>n` unfolding pos_le_divide_eq[OF `1-c>0`]
36778
739a9379e29b avoid using real-specific versions of generic lemmas
huffman
parents: 36669
diff changeset
  6542
          by (auto simp add: mult_commute dist_commute)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6543
        also have "\<dots> \<le> c ^ N * d * (1 - c ^ (m - n)) / (1 - c)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6544
          using mult_right_mono[OF * order_less_imp_le[OF **]]
36778
739a9379e29b avoid using real-specific versions of generic lemmas
huffman
parents: 36669
diff changeset
  6545
          unfolding mult_assoc by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6546
        also have "\<dots> < (e * (1 - c) / d) * d * (1 - c ^ (m - n)) / (1 - c)"
36778
739a9379e29b avoid using real-specific versions of generic lemmas
huffman
parents: 36669
diff changeset
  6547
          using mult_strict_right_mono[OF N **] unfolding mult_assoc by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6548
        also have "\<dots> = e * (1 - c ^ (m - n))" using c and `d>0` and `1 - c > 0` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6549
        also have "\<dots> \<le> e" using c and `1 - c ^ (m - n) > 0` and `e>0` using mult_right_le_one_le[of e "1 - c ^ (m - n)"] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6550
        finally have  "dist (z m) (z n) < e" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6551
      } note * = this
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6552
      { fix m n::nat assume as:"N\<le>m" "N\<le>n"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6553
        hence "dist (z n) (z m) < e"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6554
        proof(cases "n = m")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6555
          case True thus ?thesis using `e>0` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6556
        next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6557
          case False thus ?thesis using as and *[of n m] *[of m n] unfolding nat_neq_iff by (auto simp add: dist_commute)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6558
        qed }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6559
      thus ?thesis by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6560
    qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6561
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6562
  hence "Cauchy z" unfolding cauchy_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6563
  then obtain x where "x\<in>s" and x:"(z ---> x) sequentially" using s(1)[unfolded compact_def complete_def, THEN spec[where x=z]] and z_in_s by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6564
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6565
  def e \<equiv> "dist (f x) x"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6566
  have "e = 0" proof(rule ccontr)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6567
    assume "e \<noteq> 0" hence "e>0" unfolding e_def using zero_le_dist[of "f x" x]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6568
      by (metis dist_eq_0_iff dist_nz e_def)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6569
    then obtain N where N:"\<forall>n\<ge>N. dist (z n) x < e / 2"
44907
93943da0a010 remove redundant lemma Lim_sequentially in favor of lemma LIMSEQ_def
huffman
parents: 44905
diff changeset
  6570
      using x[unfolded LIMSEQ_def, THEN spec[where x="e/2"]] by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6571
    hence N':"dist (z N) x < e / 2" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6572
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6573
    have *:"c * dist (z N) x \<le> dist (z N) x" unfolding mult_le_cancel_right2
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6574
      using zero_le_dist[of "z N" x] and c
36778
739a9379e29b avoid using real-specific versions of generic lemmas
huffman
parents: 36669
diff changeset
  6575
      by (metis dist_eq_0_iff dist_nz order_less_asym less_le)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6576
    have "dist (f (z N)) (f x) \<le> c * dist (z N) x" using lipschitz[THEN bspec[where x="z N"], THEN bspec[where x=x]]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6577
      using z_in_s[of N] `x\<in>s` using c by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6578
    also have "\<dots> < e / 2" using N' and c using * by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6579
    finally show False unfolding fzn
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6580
      using N[THEN spec[where x="Suc N"]] and dist_triangle_half_r[of "z (Suc N)" "f x" e x]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6581
      unfolding e_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6582
  qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6583
  hence "f x = x" unfolding e_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6584
  moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6585
  { fix y assume "f y = y" "y\<in>s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6586
    hence "dist x y \<le> c * dist x y" using lipschitz[THEN bspec[where x=x], THEN bspec[where x=y]]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6587
      using `x\<in>s` and `f x = x` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6588
    hence "dist x y = 0" unfolding mult_le_cancel_right1
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6589
      using c and zero_le_dist[of x y] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6590
    hence "y = x" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6591
  }
34999
5312d2ffee3b Changed 'bounded unique existential quantifiers' from a constant to syntax translation.
hoelzl
parents: 34964
diff changeset
  6592
  ultimately show ?thesis using `x\<in>s` by blast+
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6593
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6594
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  6595
subsection {* Edelstein fixed point theorem *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6596
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6597
lemma edelstein_fix:
50970
3e5b67f85bf9 generalized theorem edelstein_fix to class metric_space
huffman
parents: 50955
diff changeset
  6598
  fixes s :: "'a::metric_space set"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6599
  assumes s:"compact s" "s \<noteq> {}" and gs:"(g ` s) \<subseteq> s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6600
      and dist:"\<forall>x\<in>s. \<forall>y\<in>s. x \<noteq> y \<longrightarrow> dist (g x) (g y) < dist x y"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6601
  shows "\<exists>! x\<in>s. g x = x"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6602
proof(cases "\<exists>x\<in>s. g x \<noteq> x")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6603
  obtain x where "x\<in>s" using s(2) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6604
  case False hence g:"\<forall>x\<in>s. g x = x" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6605
  { fix y assume "y\<in>s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6606
    hence "x = y" using `x\<in>s` and dist[THEN bspec[where x=x], THEN bspec[where x=y]]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6607
      unfolding g[THEN bspec[where x=x], OF `x\<in>s`]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6608
      unfolding g[THEN bspec[where x=y], OF `y\<in>s`] by auto  }
34999
5312d2ffee3b Changed 'bounded unique existential quantifiers' from a constant to syntax translation.
hoelzl
parents: 34964
diff changeset
  6609
  thus ?thesis using `x\<in>s` and g by blast+
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6610
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6611
  case True
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6612
  then obtain x where [simp]:"x\<in>s" and "g x \<noteq> x" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6613
  { fix x y assume "x \<in> s" "y \<in> s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6614
    hence "dist (g x) (g y) \<le> dist x y"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6615
      using dist[THEN bspec[where x=x], THEN bspec[where x=y]] by auto } note dist' = this
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6616
  def y \<equiv> "g x"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6617
  have [simp]:"y\<in>s" unfolding y_def using gs[unfolded image_subset_iff] and `x\<in>s` by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6618
  def f \<equiv> "\<lambda>n. g ^^ n"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6619
  have [simp]:"\<And>n z. g (f n z) = f (Suc n) z" unfolding f_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6620
  have [simp]:"\<And>z. f 0 z = z" unfolding f_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6621
  { fix n::nat and z assume "z\<in>s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6622
    have "f n z \<in> s" unfolding f_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6623
    proof(induct n)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6624
      case 0 thus ?case using `z\<in>s` by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6625
    next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6626
      case (Suc n) thus ?case using gs[unfolded image_subset_iff] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6627
    qed } note fs = this
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6628
  { fix m n ::nat assume "m\<le>n"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6629
    fix w z assume "w\<in>s" "z\<in>s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6630
    have "dist (f n w) (f n z) \<le> dist (f m w) (f m z)" using `m\<le>n`
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6631
    proof(induct n)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6632
      case 0 thus ?case by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6633
    next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6634
      case (Suc n)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6635
      thus ?case proof(cases "m\<le>n")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6636
        case True thus ?thesis using Suc(1)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6637
          using dist'[OF fs fs, OF `w\<in>s` `z\<in>s`, of n n] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6638
      next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6639
        case False hence mn:"m = Suc n" using Suc(2) by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6640
        show ?thesis unfolding mn  by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6641
      qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6642
    qed } note distf = this
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6643
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6644
  def h \<equiv> "\<lambda>n. (f n x, f n y)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6645
  let ?s2 = "s \<times> s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6646
  obtain l r where "l\<in>?s2" and r:"subseq r" and lr:"((h \<circ> r) ---> l) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6647
    using compact_Times [OF s(1) s(1), unfolded compact_def, THEN spec[where x=h]] unfolding  h_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6648
    using fs[OF `x\<in>s`] and fs[OF `y\<in>s`] by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6649
  def a \<equiv> "fst l" def b \<equiv> "snd l"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6650
  have lab:"l = (a, b)" unfolding a_def b_def by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6651
  have [simp]:"a\<in>s" "b\<in>s" unfolding a_def b_def using `l\<in>?s2` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6652
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6653
  have lima:"((fst \<circ> (h \<circ> r)) ---> a) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6654
   and limb:"((snd \<circ> (h \<circ> r)) ---> b) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6655
    using lr
44167
e81d676d598e avoid duplicate rule warnings
huffman
parents: 44139
diff changeset
  6656
    unfolding o_def a_def b_def by (rule tendsto_intros)+
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6657
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6658
  { fix n::nat
50970
3e5b67f85bf9 generalized theorem edelstein_fix to class metric_space
huffman
parents: 50955
diff changeset
  6659
    have *:"\<And>fx fy (x::'a) y. dist fx fy \<le> dist x y \<Longrightarrow> \<not> (\<bar>dist fx fy - dist a b\<bar> < dist a b - dist x y)" by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6660
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6661
    { assume as:"dist a b > dist (f n x) (f n y)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6662
      then obtain Na Nb where "\<forall>m\<ge>Na. dist (f (r m) x) a < (dist a b - dist (f n x) (f n y)) / 2"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6663
        and "\<forall>m\<ge>Nb. dist (f (r m) y) b < (dist a b - dist (f n x) (f n y)) / 2"
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46887
diff changeset
  6664
        using lima limb unfolding h_def LIMSEQ_def by (fastforce simp del: less_divide_eq_numeral1)
50970
3e5b67f85bf9 generalized theorem edelstein_fix to class metric_space
huffman
parents: 50955
diff changeset
  6665
      hence "\<bar>dist (f (r (Na + Nb + n)) x) (f (r (Na + Nb + n)) y) - dist a b\<bar> < dist a b - dist (f n x) (f n y)"
3e5b67f85bf9 generalized theorem edelstein_fix to class metric_space
huffman
parents: 50955
diff changeset
  6666
        apply -
3e5b67f85bf9 generalized theorem edelstein_fix to class metric_space
huffman
parents: 50955
diff changeset
  6667
        apply (drule_tac x="Na+Nb+n" in spec, drule mp, simp)
3e5b67f85bf9 generalized theorem edelstein_fix to class metric_space
huffman
parents: 50955
diff changeset
  6668
        apply (drule_tac x="Na+Nb+n" in spec, drule mp, simp)
3e5b67f85bf9 generalized theorem edelstein_fix to class metric_space
huffman
parents: 50955
diff changeset
  6669
        apply (drule (1) add_strict_mono, simp only: real_sum_of_halves)
3e5b67f85bf9 generalized theorem edelstein_fix to class metric_space
huffman
parents: 50955
diff changeset
  6670
        apply (erule le_less_trans [rotated])
3e5b67f85bf9 generalized theorem edelstein_fix to class metric_space
huffman
parents: 50955
diff changeset
  6671
        apply (erule thin_rl)
3e5b67f85bf9 generalized theorem edelstein_fix to class metric_space
huffman
parents: 50955
diff changeset
  6672
        apply (rule abs_leI)
3e5b67f85bf9 generalized theorem edelstein_fix to class metric_space
huffman
parents: 50955
diff changeset
  6673
        apply (simp add: diff_le_iff add_assoc)
3e5b67f85bf9 generalized theorem edelstein_fix to class metric_space
huffman
parents: 50955
diff changeset
  6674
        apply (rule order_trans [OF dist_triangle add_left_mono])
3e5b67f85bf9 generalized theorem edelstein_fix to class metric_space
huffman
parents: 50955
diff changeset
  6675
        apply (subst add_commute, rule dist_triangle2)
3e5b67f85bf9 generalized theorem edelstein_fix to class metric_space
huffman
parents: 50955
diff changeset
  6676
        apply (simp add: diff_le_iff add_assoc)
3e5b67f85bf9 generalized theorem edelstein_fix to class metric_space
huffman
parents: 50955
diff changeset
  6677
        apply (rule order_trans [OF dist_triangle3 add_left_mono])
3e5b67f85bf9 generalized theorem edelstein_fix to class metric_space
huffman
parents: 50955
diff changeset
  6678
        apply (subst add_commute, rule dist_triangle)
3e5b67f85bf9 generalized theorem edelstein_fix to class metric_space
huffman
parents: 50955
diff changeset
  6679
        done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6680
      moreover
50970
3e5b67f85bf9 generalized theorem edelstein_fix to class metric_space
huffman
parents: 50955
diff changeset
  6681
      have "\<bar>dist (f (r (Na + Nb + n)) x) (f (r (Na + Nb + n)) y) - dist a b\<bar> \<ge> dist a b - dist (f n x) (f n y)"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6682
        using distf[of n "r (Na+Nb+n)", OF _ `x\<in>s` `y\<in>s`]
50937
d249ef928ae1 removed subseq_bigger (replaced by seq_suble)
hoelzl
parents: 50936
diff changeset
  6683
        using seq_suble[OF r, of "Na+Nb+n"]
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6684
        using *[of "f (r (Na + Nb + n)) x" "f (r (Na + Nb + n)) y" "f n x" "f n y"] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6685
      ultimately have False by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6686
    }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6687
    hence "dist a b \<le> dist (f n x) (f n y)" by(rule ccontr)auto }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6688
  note ab_fn = this
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6689
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6690
  have [simp]:"a = b" proof(rule ccontr)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6691
    def e \<equiv> "dist a b - dist (g a) (g b)"
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 44668
diff changeset
  6692
    assume "a\<noteq>b" hence "e > 0" unfolding e_def using dist by fastforce
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6693
    hence "\<exists>n. dist (f n x) a < e/2 \<and> dist (f n y) b < e/2"
44907
93943da0a010 remove redundant lemma Lim_sequentially in favor of lemma LIMSEQ_def
huffman
parents: 44905
diff changeset
  6694
      using lima limb unfolding LIMSEQ_def
93943da0a010 remove redundant lemma Lim_sequentially in favor of lemma LIMSEQ_def
huffman
parents: 44905
diff changeset
  6695
      apply (auto elim!: allE[where x="e/2"]) apply(rename_tac N N', rule_tac x="r (max N N')" in exI) unfolding h_def by fastforce
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6696
    then obtain n where n:"dist (f n x) a < e/2 \<and> dist (f n y) b < e/2" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6697
    have "dist (f (Suc n) x) (g a) \<le> dist (f n x) a"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6698
      using dist[THEN bspec[where x="f n x"], THEN bspec[where x="a"]] and fs by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6699
    moreover have "dist (f (Suc n) y) (g b) \<le> dist (f n y) b"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6700
      using dist[THEN bspec[where x="f n y"], THEN bspec[where x="b"]] and fs by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6701
    ultimately have "dist (f (Suc n) x) (g a) + dist (f (Suc n) y) (g b) < e" using n by auto
50970
3e5b67f85bf9 generalized theorem edelstein_fix to class metric_space
huffman
parents: 50955
diff changeset
  6702
    thus False unfolding e_def using ab_fn[of "Suc n"]
3e5b67f85bf9 generalized theorem edelstein_fix to class metric_space
huffman
parents: 50955
diff changeset
  6703
      using dist_triangle2 [of "f (Suc n) y" "g a" "g b"]
3e5b67f85bf9 generalized theorem edelstein_fix to class metric_space
huffman
parents: 50955
diff changeset
  6704
      using dist_triangle2 [of "f (Suc n) x" "f (Suc n) y" "g a"]
3e5b67f85bf9 generalized theorem edelstein_fix to class metric_space
huffman
parents: 50955
diff changeset
  6705
      by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6706
  qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6707
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6708
  have [simp]:"\<And>n. f (Suc n) x = f n y" unfolding f_def y_def by(induct_tac n)auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6709
  { fix x y assume "x\<in>s" "y\<in>s" moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6710
    fix e::real assume "e>0" ultimately
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 44668
diff changeset
  6711
    have "dist y x < e \<longrightarrow> dist (g y) (g x) < e" using dist by fastforce }
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  6712
  hence "continuous_on s g" unfolding continuous_on_iff by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6713
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6714
  hence "((snd \<circ> h \<circ> r) ---> g a) sequentially" unfolding continuous_on_sequentially
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6715
    apply (rule allE[where x="\<lambda>n. (fst \<circ> h \<circ> r) n"]) apply (erule ballE[where x=a])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6716
    using lima unfolding h_def o_def using fs[OF `x\<in>s`] by (auto simp add: y_def)
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41969
diff changeset
  6717
  hence "g a = a" using tendsto_unique[OF trivial_limit_sequentially limb, of "g a"]
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6718
    unfolding `a=b` and o_assoc by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6719
  moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6720
  { fix x assume "x\<in>s" "g x = x" "x\<noteq>a"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6721
    hence "False" using dist[THEN bspec[where x=a], THEN bspec[where x=x]]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6722
      using `g a = a` and `a\<in>s` by auto  }
34999
5312d2ffee3b Changed 'bounded unique existential quantifiers' from a constant to syntax translation.
hoelzl
parents: 34964
diff changeset
  6723
  ultimately show "\<exists>!x\<in>s. g x = x" using `a\<in>s` by blast
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6724
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6725
44131
5fc334b94e00 declare tendsto_const [intro] (accidentally removed in 230a8665c919)
huffman
parents: 44129
diff changeset
  6726
declare tendsto_const [intro] (* FIXME: move *)
5fc334b94e00 declare tendsto_const [intro] (accidentally removed in 230a8665c919)
huffman
parents: 44129
diff changeset
  6727
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6728
end