src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy
author huffman
Thu, 17 Jan 2013 15:28:53 -0800
changeset 50970 3e5b67f85bf9
parent 50955 ada575c605e1
child 50971 5e3d3d690975
permissions -rw-r--r--
generalized theorem edelstein_fix to class metric_space
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
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    12
  "~~/src/HOL/Library/Diagonal_Subsequence"
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
    13
  "~~/src/HOL/Library/Countable_Set"
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    14
  "~~/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
    15
  "~~/src/HOL/Library/FuncSet"
50938
hoelzl
parents: 50937
diff changeset
    16
  Linear_Algebra
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    17
  Norm_Arith
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    18
begin
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    19
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
    20
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
    21
  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
    22
50942
1aa1a7991fd9 move auxiliary lemma to top
hoelzl
parents: 50941
diff changeset
    23
(* TODO: Move this to RComplete.thy -- would need to include Glb into RComplete *)
1aa1a7991fd9 move auxiliary lemma to top
hoelzl
parents: 50941
diff changeset
    24
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
    25
  apply (frule isGlb_isLb)
1aa1a7991fd9 move auxiliary lemma to top
hoelzl
parents: 50941
diff changeset
    26
  apply (frule_tac x = y in isGlb_isLb)
1aa1a7991fd9 move auxiliary lemma to top
hoelzl
parents: 50941
diff changeset
    27
  apply (blast intro!: order_antisym dest!: isGlb_le_isLb)
1aa1a7991fd9 move auxiliary lemma to top
hoelzl
parents: 50941
diff changeset
    28
  done
1aa1a7991fd9 move auxiliary lemma to top
hoelzl
parents: 50941
diff changeset
    29
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
    30
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
    31
  "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
    32
  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
    33
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    34
subsection {* Topological Basis *}
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    35
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    36
context topological_space
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    37
begin
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    38
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    39
definition "topological_basis B =
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    40
  ((\<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
    41
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    42
lemma topological_basis_iff:
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    43
  assumes "\<And>B'. B' \<in> B \<Longrightarrow> open B'"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    44
  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
    45
    (is "_ \<longleftrightarrow> ?rhs")
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    46
proof safe
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    47
  fix O' and x::'a
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    48
  assume H: "topological_basis B" "open O'" "x \<in> O'"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    49
  hence "(\<exists>B'\<subseteq>B. \<Union>B' = O')" by (simp add: topological_basis_def)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    50
  then obtain B' where "B' \<subseteq> B" "O' = \<Union>B'" by auto
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    51
  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
    52
next
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    53
  assume H: ?rhs
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    54
  show "topological_basis B" using assms unfolding topological_basis_def
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    55
  proof safe
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    56
    fix O'::"'a set" assume "open O'"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    57
    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
    58
      by (force intro: bchoice simp: Bex_def)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    59
    thus "\<exists>B'\<subseteq>B. \<Union>B' = O'"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    60
      by (auto intro: exI[where x="{f x |x. x \<in> O'}"])
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    61
  qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    62
qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    63
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    64
lemma topological_basisI:
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    65
  assumes "\<And>B'. B' \<in> B \<Longrightarrow> open B'"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    66
  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
    67
  shows "topological_basis B"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    68
  using assms by (subst topological_basis_iff) auto
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    69
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    70
lemma topological_basisE:
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    71
  fixes O'
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    72
  assumes "topological_basis B"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    73
  assumes "open O'"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    74
  assumes "x \<in> O'"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    75
  obtains B' where "B' \<in> B" "x \<in> B'" "B' \<subseteq> O'"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    76
proof atomize_elim
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    77
  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
    78
  with topological_basis_iff assms
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    79
  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
    80
qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    81
50094
84ddcf5364b4 allow arbitrary enumerations of basis in locale for generation of borel sets
immler
parents: 50087
diff changeset
    82
lemma topological_basis_open:
84ddcf5364b4 allow arbitrary enumerations of basis in locale for generation of borel sets
immler
parents: 50087
diff changeset
    83
  assumes "topological_basis B"
84ddcf5364b4 allow arbitrary enumerations of basis in locale for generation of borel sets
immler
parents: 50087
diff changeset
    84
  assumes "X \<in> B"
84ddcf5364b4 allow arbitrary enumerations of basis in locale for generation of borel sets
immler
parents: 50087
diff changeset
    85
  shows "open X"
84ddcf5364b4 allow arbitrary enumerations of basis in locale for generation of borel sets
immler
parents: 50087
diff changeset
    86
  using assms
84ddcf5364b4 allow arbitrary enumerations of basis in locale for generation of borel sets
immler
parents: 50087
diff changeset
    87
  by (simp add: topological_basis_def)
84ddcf5364b4 allow arbitrary enumerations of basis in locale for generation of borel sets
immler
parents: 50087
diff changeset
    88
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
    89
lemma basis_dense:
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
    90
  fixes B::"'a set set" and f::"'a set \<Rightarrow> 'a"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
    91
  assumes "topological_basis B"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
    92
  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
    93
  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
    94
proof (intro allI impI)
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
    95
  fix X::"'a set" assume "open X" "X \<noteq> {}"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
    96
  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
    97
  guess B' . note B' = this
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
    98
  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
    99
qed
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   100
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   101
end
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   102
50882
a382bf90867e move prod instantiation of second_countable_topology to its definition
hoelzl
parents: 50881
diff changeset
   103
lemma topological_basis_prod:
a382bf90867e move prod instantiation of second_countable_topology to its definition
hoelzl
parents: 50881
diff changeset
   104
  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
   105
  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
   106
  unfolding topological_basis_def
a382bf90867e move prod instantiation of second_countable_topology to its definition
hoelzl
parents: 50881
diff changeset
   107
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
   108
  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
   109
  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
   110
  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
   111
    fix x y assume "(x, y) \<in> S"
a382bf90867e move prod instantiation of second_countable_topology to its definition
hoelzl
parents: 50881
diff changeset
   112
    from open_prod_elim[OF `open S` this]
a382bf90867e move prod instantiation of second_countable_topology to its definition
hoelzl
parents: 50881
diff changeset
   113
    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
   114
      by (metis mem_Sigma_iff)
a382bf90867e move prod instantiation of second_countable_topology to its definition
hoelzl
parents: 50881
diff changeset
   115
    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
   116
    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
   117
    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
   118
      by (intro UN_I[of "(A0, B0)"]) auto
a382bf90867e move prod instantiation of second_countable_topology to its definition
hoelzl
parents: 50881
diff changeset
   119
  qed auto
a382bf90867e move prod instantiation of second_countable_topology to its definition
hoelzl
parents: 50881
diff changeset
   120
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
   121
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   122
subsection {* Countable Basis *}
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   123
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   124
locale countable_basis =
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   125
  fixes B::"'a::topological_space set set"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   126
  assumes is_basis: "topological_basis B"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   127
  assumes countable_basis: "countable B"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   128
begin
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   129
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   130
lemma open_countable_basis_ex:
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   131
  assumes "open X"
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   132
  shows "\<exists>B' \<subseteq> B. X = Union B'"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   133
  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
   134
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   135
lemma open_countable_basisE:
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   136
  assumes "open X"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   137
  obtains B' where "B' \<subseteq> B" "X = Union B'"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   138
  using assms open_countable_basis_ex by (atomize_elim) simp
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   139
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   140
lemma countable_dense_exists:
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   141
  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
   142
proof -
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   143
  let ?f = "(\<lambda>B'. SOME x. x \<in> B')"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   144
  have "countable (?f ` B)" using countable_basis by simp
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   145
  with basis_dense[OF is_basis, of ?f] show ?thesis
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   146
    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
   147
qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   148
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   149
lemma countable_dense_setE:
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   150
  obtains D :: "'a set"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   151
  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
   152
  using countable_dense_exists by blast
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   153
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   154
text {* Construction of an increasing sequence approximating open sets,
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   155
  therefore basis which is closed under union. *}
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   156
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   157
definition union_closed_basis::"'a set set" where
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   158
  "union_closed_basis = (\<lambda>l. \<Union>set l) ` lists B"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   159
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   160
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
   161
proof (rule topological_basisI)
84ddcf5364b4 allow arbitrary enumerations of basis in locale for generation of borel sets
immler
parents: 50087
diff changeset
   162
  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
   163
  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
   164
  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
   165
    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
   166
next
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   167
  fix B' assume "B' \<in> union_closed_basis"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   168
  thus "open B'"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   169
    using topological_basis_open[OF is_basis]
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   170
    by (auto simp: union_closed_basis_def)
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   171
qed
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   172
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   173
lemma countable_union_closed_basis: "countable union_closed_basis"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   174
  unfolding union_closed_basis_def using countable_basis by simp
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   175
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   176
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
   177
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   178
lemma union_closed_basis_ex:
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   179
 assumes X: "X \<in> union_closed_basis"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   180
 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
   181
proof -
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   182
  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
   183
  thus ?thesis by auto
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   184
qed
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_basisE:
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   187
  assumes "X \<in> union_closed_basis"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   188
  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
   189
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   190
lemma union_closed_basisI:
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   191
  assumes "finite B'" "X = \<Union>B'" "B' \<subseteq> B"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   192
  shows "X \<in> union_closed_basis"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   193
proof -
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   194
  from finite_list[OF `finite B'`] guess l ..
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   195
  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
   196
qed
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 empty_basisI[intro]: "{} \<in> union_closed_basis"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   199
  by (rule union_closed_basisI[of "{}"]) auto
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   200
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   201
lemma union_basisI[intro]:
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   202
  assumes "X \<in> union_closed_basis" "Y \<in> union_closed_basis"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   203
  shows "X \<union> Y \<in> union_closed_basis"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   204
  using assms by (auto intro: union_closed_basisI elim!:union_closed_basisE)
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   205
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   206
lemma open_imp_Union_of_incseq:
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   207
  assumes "open X"
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   208
  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
   209
proof -
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   210
  from open_countable_basis_ex[OF `open X`]
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   211
  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
   212
  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
   213
  show ?thesis
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   214
  proof cases
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   215
    assume "B' \<noteq> {}"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   216
    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
   217
    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
   218
    have "incseq S" by (force simp: S_def incseq_Suc_iff)
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   219
    moreover
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   220
    have "(\<Union>j. S j) = X" unfolding B'
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   221
    proof safe
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   222
      fix x X assume "X \<in> B'" "x \<in> X"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   223
      then obtain n where "X = from_nat_into B' n"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   224
        by (metis `countable B'` from_nat_into_surj)
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   225
      also have "\<dots> \<subseteq> S n" by (auto simp: S_def)
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   226
      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
   227
    next
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   228
      fix x n
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   229
      assume "x \<in> S n"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   230
      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
   231
        by (simp add: S_def)
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   232
      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
   233
      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
   234
      finally show "x \<in> \<Union>B'" .
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   235
    qed
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   236
    moreover have "range S \<subseteq> union_closed_basis" using B'
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   237
      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
   238
    ultimately show ?thesis by auto
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   239
  qed (auto simp: B')
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   240
qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   241
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   242
lemma open_incseqE:
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   243
  assumes "open X"
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   244
  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
   245
  using open_imp_Union_of_incseq assms by atomize_elim
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   246
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   247
end
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   248
50883
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   249
class first_countable_topology = topological_space +
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   250
  assumes first_countable_basis:
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   251
    "\<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
   252
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   253
lemma (in first_countable_topology) countable_basis_at_decseq:
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   254
  obtains A :: "nat \<Rightarrow> 'a set" where
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   255
    "\<And>i. open (A i)" "\<And>i. x \<in> (A i)"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   256
    "\<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
   257
proof atomize_elim
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   258
  from first_countable_basis[of x] obtain A
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   259
    where "countable A"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   260
    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
   261
    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
   262
  then have "A \<noteq> {}" by auto
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   263
  with `countable A` have r: "A = range (from_nat_into A)" by auto
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   264
  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
   265
  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
   266
      (\<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
   267
  proof (safe intro!: exI[of _ F])
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   268
    fix i
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   269
    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
   270
    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
   271
  next
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   272
    fix S assume "open S" "x \<in> S"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   273
    from incl[OF this] obtain i where "F i \<subseteq> S"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   274
      by (subst (asm) r) (auto simp: F_def)
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   275
    moreover have "\<And>j. i \<le> j \<Longrightarrow> F j \<subseteq> F i"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   276
      by (auto simp: F_def)
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   277
    ultimately show "eventually (\<lambda>i. F i \<subseteq> S) sequentially"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   278
      by (auto simp: eventually_sequentially)
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   279
  qed
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   280
qed
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   281
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   282
lemma (in first_countable_topology) first_countable_basisE:
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   283
  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
   284
    "\<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
   285
  using first_countable_basis[of x]
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   286
  by atomize_elim auto
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   287
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   288
instance prod :: (first_countable_topology, first_countable_topology) first_countable_topology
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   289
proof
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   290
  fix x :: "'a \<times> 'b"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   291
  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
   292
  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
   293
  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
   294
  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
   295
    fix a b assume x: "a \<in> A" "b \<in> B"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   296
    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
   297
      unfolding mem_Times_iff by (auto intro: open_Times)
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   298
  next
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   299
    fix S assume "open S" "x \<in> S"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   300
    from open_prod_elim[OF this] guess a' b' .
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   301
    moreover with A(4)[of a'] B(4)[of b']
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   302
    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
   303
    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
   304
      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
   305
  qed (simp add: A B)
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   306
qed
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   307
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   308
instance metric_space \<subseteq> first_countable_topology
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   309
proof
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   310
  fix x :: 'a
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   311
  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
   312
  proof (intro exI, safe)
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   313
    fix S assume "open S" "x \<in> S"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   314
    then obtain r where "0 < r" "{y. dist x y < r} \<subseteq> S"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   315
      by (auto simp: open_dist dist_commute subset_eq)
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   316
    moreover from reals_Archimedean[OF `0 < r`] guess n ..
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   317
    moreover
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   318
    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
   319
      by (auto simp: inverse_eq_divide)
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   320
    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
   321
      by auto
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   322
  qed (auto intro: open_ball)
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   323
qed
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   324
50881
ae630bab13da renamed countable_basis_space to second_countable_topology
hoelzl
parents: 50526
diff changeset
   325
class second_countable_topology = topological_space +
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   326
  assumes ex_countable_basis:
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   327
    "\<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
   328
50881
ae630bab13da renamed countable_basis_space to second_countable_topology
hoelzl
parents: 50526
diff changeset
   329
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
   330
  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
   331
50882
a382bf90867e move prod instantiation of second_countable_topology to its definition
hoelzl
parents: 50881
diff changeset
   332
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
   333
proof
a382bf90867e move prod instantiation of second_countable_topology to its definition
hoelzl
parents: 50881
diff changeset
   334
  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
   335
    using ex_countable_basis by auto
a382bf90867e move prod instantiation of second_countable_topology to its definition
hoelzl
parents: 50881
diff changeset
   336
  moreover
a382bf90867e move prod instantiation of second_countable_topology to its definition
hoelzl
parents: 50881
diff changeset
   337
  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
   338
    using ex_countable_basis by auto
a382bf90867e move prod instantiation of second_countable_topology to its definition
hoelzl
parents: 50881
diff changeset
   339
  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
   340
    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
   341
qed
a382bf90867e move prod instantiation of second_countable_topology to its definition
hoelzl
parents: 50881
diff changeset
   342
50883
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   343
instance second_countable_topology \<subseteq> first_countable_topology
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   344
proof
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   345
  fix x :: 'a
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   346
  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
   347
  then have B: "countable B" "topological_basis B"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   348
    using countable_basis is_basis
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   349
    by (auto simp: countable_basis is_basis)
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   350
  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
   351
    by (intro exI[of _ "{b\<in>B. x \<in> b}"])
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   352
       (fastforce simp: topological_space_class.topological_basis_def)
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   353
qed
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   354
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   355
subsection {* Polish spaces *}
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   356
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   357
text {* Textbooks define Polish spaces as completely metrizable.
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   358
  We assume the topology to be complete for a given metric. *}
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   359
50881
ae630bab13da renamed countable_basis_space to second_countable_topology
hoelzl
parents: 50526
diff changeset
   360
class polish_space = complete_space + second_countable_topology
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   361
44517
68e8eb0ce8aa minimize imports
huffman
parents: 44516
diff changeset
   362
subsection {* General notion of a topology as a value *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   363
44170
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   364
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
   365
typedef 'a topology = "{L::('a set) \<Rightarrow> bool. istopology L}"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   366
  morphisms "openin" "topology"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   367
  unfolding istopology_def by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   368
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   369
lemma istopology_open_in[intro]: "istopology(openin U)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   370
  using openin[of U] by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   371
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   372
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
   373
  using topology_inverse[unfolded mem_Collect_eq] .
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   374
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   375
lemma topology_inverse_iff: "istopology U \<longleftrightarrow> openin (topology U) = U"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   376
  using topology_inverse[of U] istopology_open_in[of "topology U"] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   377
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   378
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
   379
proof-
49711
e5aaae7eadc9 tuned proofs;
wenzelm
parents: 48125
diff changeset
   380
  { assume "T1=T2"
e5aaae7eadc9 tuned proofs;
wenzelm
parents: 48125
diff changeset
   381
    hence "\<forall>S. openin T1 S \<longleftrightarrow> openin T2 S" by simp }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   382
  moreover
49711
e5aaae7eadc9 tuned proofs;
wenzelm
parents: 48125
diff changeset
   383
  { 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
   384
    hence "openin T1 = openin T2" by (simp add: fun_eq_iff)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   385
    hence "topology (openin T1) = topology (openin T2)" by simp
49711
e5aaae7eadc9 tuned proofs;
wenzelm
parents: 48125
diff changeset
   386
    hence "T1 = T2" unfolding openin_inverse .
e5aaae7eadc9 tuned proofs;
wenzelm
parents: 48125
diff changeset
   387
  }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   388
  ultimately show ?thesis by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   389
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   390
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   391
text{* Infer the "universe" from union of all sets in the topology. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   392
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   393
definition "topspace T =  \<Union>{S. openin T S}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   394
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   395
subsubsection {* Main properties of open sets *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   396
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   397
lemma openin_clauses:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   398
  fixes U :: "'a topology"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   399
  shows "openin U {}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   400
  "\<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
   401
  "\<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
   402
  using openin[of U] unfolding istopology_def mem_Collect_eq
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   403
  by fast+
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_subset[intro]: "openin U S \<Longrightarrow> S \<subseteq> topspace U"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   406
  unfolding topspace_def by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   407
lemma openin_empty[simp]: "openin U {}" by (simp add: openin_clauses)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   408
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   409
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
   410
  using openin_clauses by simp
06475a1547cb fix lots of looping simp calls and other warnings
huffman
parents: 36360
diff changeset
   411
06475a1547cb fix lots of looping simp calls and other warnings
huffman
parents: 36360
diff changeset
   412
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
   413
  using openin_clauses by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   414
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   415
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
   416
  using openin_Union[of "{S,T}" U] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   417
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   418
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
   419
49711
e5aaae7eadc9 tuned proofs;
wenzelm
parents: 48125
diff changeset
   420
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
   421
  (is "?lhs \<longleftrightarrow> ?rhs")
36584
1535841fc2e9 prove lemma openin_subopen without using choice
huffman
parents: 36442
diff changeset
   422
proof
49711
e5aaae7eadc9 tuned proofs;
wenzelm
parents: 48125
diff changeset
   423
  assume ?lhs
e5aaae7eadc9 tuned proofs;
wenzelm
parents: 48125
diff changeset
   424
  then show ?rhs by auto
36584
1535841fc2e9 prove lemma openin_subopen without using choice
huffman
parents: 36442
diff changeset
   425
next
1535841fc2e9 prove lemma openin_subopen without using choice
huffman
parents: 36442
diff changeset
   426
  assume H: ?rhs
1535841fc2e9 prove lemma openin_subopen without using choice
huffman
parents: 36442
diff changeset
   427
  let ?t = "\<Union>{T. openin U T \<and> T \<subseteq> S}"
1535841fc2e9 prove lemma openin_subopen without using choice
huffman
parents: 36442
diff changeset
   428
  have "openin U ?t" by (simp add: openin_Union)
1535841fc2e9 prove lemma openin_subopen without using choice
huffman
parents: 36442
diff changeset
   429
  also have "?t = S" using H by auto
1535841fc2e9 prove lemma openin_subopen without using choice
huffman
parents: 36442
diff changeset
   430
  finally show "openin U S" .
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   431
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   432
49711
e5aaae7eadc9 tuned proofs;
wenzelm
parents: 48125
diff changeset
   433
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   434
subsubsection {* Closed sets *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   435
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   436
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
   437
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   438
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
   439
lemma closedin_empty[simp]: "closedin U {}" by (simp add: closedin_def)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   440
lemma closedin_topspace[intro,simp]:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   441
  "closedin U (topspace U)" by (simp add: closedin_def)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   442
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
   443
  by (auto simp add: Diff_Un closedin_def)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   444
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   445
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
   446
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
   447
  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
   448
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   449
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
   450
  using closedin_Inter[of "{S,T}" U] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   451
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   452
lemma Diff_Diff_Int: "A - (A - B) = A \<inter> B" by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   453
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
   454
  apply (auto simp add: closedin_def Diff_Diff_Int inf_absorb2)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   455
  apply (metis openin_subset subset_eq)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   456
  done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   457
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   458
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
   459
  by (simp add: openin_closedin_eq)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   460
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   461
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
   462
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   463
  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
   464
    by (auto simp add: topspace_def openin_subset)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   465
  then show ?thesis using oS cT by (auto simp add: closedin_def)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   466
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   467
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   468
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
   469
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   470
  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
   471
    by (auto simp add: topspace_def )
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   472
  then show ?thesis using oS cT by (auto simp add: openin_closedin_eq)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   473
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   474
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   475
subsubsection {* Subspace topology *}
44170
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   476
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   477
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
   478
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   479
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
   480
  (is "istopology ?L")
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   481
proof-
44170
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   482
  have "?L {}" by blast
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   483
  {fix A B assume A: "?L A" and B: "?L B"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   484
    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
   485
    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
   486
    then have "?L (A \<inter> B)" by blast}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   487
  moreover
44170
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   488
  {fix K assume K: "K \<subseteq> Collect ?L"
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   489
    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
   490
      apply (rule set_eqI)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   491
      apply (simp add: Ball_def image_iff)
44170
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   492
      by metis
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   493
    from K[unfolded th0 subset_image_iff]
44170
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   494
    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
   495
    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
   496
    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
   497
    ultimately have "?L (\<Union>K)" by blast}
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   498
  ultimately show ?thesis
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   499
    unfolding subset_eq mem_Collect_eq istopology_def by blast
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   500
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   501
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   502
lemma openin_subtopology:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   503
  "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
   504
  unfolding subtopology_def topology_inverse'[OF istopology_subtopology]
44170
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   505
  by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   506
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   507
lemma topspace_subtopology: "topspace(subtopology U V) = topspace U \<inter> V"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   508
  by (auto simp add: topspace_def openin_subtopology)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   509
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   510
lemma closedin_subtopology:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   511
  "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
   512
  unfolding closedin_def topspace_subtopology
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   513
  apply (simp add: openin_subtopology)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   514
  apply (rule iffI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   515
  apply clarify
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   516
  apply (rule_tac x="topspace U - T" in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   517
  by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   518
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   519
lemma openin_subtopology_refl: "openin (subtopology U V) V \<longleftrightarrow> V \<subseteq> topspace U"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   520
  unfolding openin_subtopology
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   521
  apply (rule iffI, clarify)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   522
  apply (frule openin_subset[of U])  apply blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   523
  apply (rule exI[where x="topspace U"])
49711
e5aaae7eadc9 tuned proofs;
wenzelm
parents: 48125
diff changeset
   524
  apply auto
e5aaae7eadc9 tuned proofs;
wenzelm
parents: 48125
diff changeset
   525
  done
e5aaae7eadc9 tuned proofs;
wenzelm
parents: 48125
diff changeset
   526
e5aaae7eadc9 tuned proofs;
wenzelm
parents: 48125
diff changeset
   527
lemma subtopology_superset:
e5aaae7eadc9 tuned proofs;
wenzelm
parents: 48125
diff changeset
   528
  assumes UV: "topspace U \<subseteq> V"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   529
  shows "subtopology U V = U"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   530
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   531
  {fix S
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   532
    {fix T assume T: "openin U T" "S = T \<inter> V"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   533
      from T openin_subset[OF T(1)] UV have eq: "S = T" by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   534
      have "openin U S" unfolding eq using T by blast}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   535
    moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   536
    {assume S: "openin U S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   537
      hence "\<exists>T. openin U T \<and> S = T \<inter> V"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   538
        using openin_subset[OF S] UV by auto}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   539
    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
   540
  then show ?thesis unfolding topology_eq openin_subtopology by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   541
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   542
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   543
lemma subtopology_topspace[simp]: "subtopology U (topspace U) = U"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   544
  by (simp add: subtopology_superset)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   545
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   546
lemma subtopology_UNIV[simp]: "subtopology U UNIV = U"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   547
  by (simp add: subtopology_superset)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   548
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   549
subsubsection {* The standard Euclidean topology *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   550
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   551
definition
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   552
  euclidean :: "'a::topological_space topology" where
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   553
  "euclidean = topology open"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   554
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   555
lemma open_openin: "open S \<longleftrightarrow> openin euclidean S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   556
  unfolding euclidean_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   557
  apply (rule cong[where x=S and y=S])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   558
  apply (rule topology_inverse[symmetric])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   559
  apply (auto simp add: istopology_def)
44170
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   560
  done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   561
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   562
lemma topspace_euclidean: "topspace euclidean = UNIV"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   563
  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
   564
  apply (rule set_eqI)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   565
  by (auto simp add: open_openin[symmetric])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   566
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   567
lemma topspace_euclidean_subtopology[simp]: "topspace (subtopology euclidean S) = S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   568
  by (simp add: topspace_euclidean topspace_subtopology)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   569
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   570
lemma closed_closedin: "closed S \<longleftrightarrow> closedin euclidean S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   571
  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
   572
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   573
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
   574
  by (simp add: open_openin openin_subopen[symmetric])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   575
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   576
text {* Basic "localization" results are handy for connectedness. *}
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   577
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   578
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
   579
  by (auto simp add: openin_subtopology open_openin[symmetric])
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   580
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   581
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
   582
  by (auto simp add: openin_open)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   583
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   584
lemma open_openin_trans[trans]:
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   585
 "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
   586
  by (metis Int_absorb1  openin_open_Int)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   587
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   588
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
   589
  by (auto simp add: openin_open)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   590
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   591
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
   592
  by (simp add: closedin_subtopology closed_closedin Int_ac)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   593
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   594
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
   595
  by (metis closedin_closed)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   596
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   597
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
   598
  apply (subgoal_tac "S \<inter> T = T" )
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   599
  apply auto
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   600
  apply (frule closedin_closed_Int[of T S])
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   601
  by simp
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   602
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   603
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
   604
  by (auto simp add: closedin_closed)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   605
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   606
lemma openin_euclidean_subtopology_iff:
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   607
  fixes S U :: "'a::metric_space set"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   608
  shows "openin (subtopology euclidean U) S
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   609
  \<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
   610
proof
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   611
  assume ?lhs thus ?rhs unfolding openin_open open_dist by blast
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   612
next
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   613
  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
   614
  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
   615
    unfolding T_def
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   616
    apply clarsimp
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   617
    apply (rule_tac x="d - dist x a" in exI)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   618
    apply (clarsimp simp add: less_diff_eq)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   619
    apply (erule rev_bexI)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   620
    apply (rule_tac x=d in exI, clarify)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   621
    apply (erule le_less_trans [OF dist_triangle])
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   622
    done
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   623
  assume ?rhs hence 2: "S = U \<inter> T"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   624
    unfolding T_def
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   625
    apply auto
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   626
    apply (drule (1) bspec, erule rev_bexI)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   627
    apply auto
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   628
    done
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   629
  from 1 2 show ?lhs
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   630
    unfolding openin_open open_dist by fast
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   631
qed
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   632
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   633
text {* These "transitivity" results are handy too *}
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   634
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   635
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
   636
  \<Longrightarrow> openin (subtopology euclidean U) S"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   637
  unfolding open_openin openin_open by blast
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   638
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   639
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
   640
  by (auto simp add: openin_open intro: openin_trans)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   641
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   642
lemma closedin_trans[trans]:
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   643
 "closedin (subtopology euclidean T) S \<Longrightarrow>
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   644
           closedin (subtopology euclidean U) T
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   645
           ==> closedin (subtopology euclidean U) S"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   646
  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
   647
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   648
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
   649
  by (auto simp add: closedin_closed intro: closedin_trans)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   650
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   651
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   652
subsection {* Open and closed balls *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   653
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   654
definition
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   655
  ball :: "'a::metric_space \<Rightarrow> real \<Rightarrow> 'a set" where
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   656
  "ball x e = {y. dist x y < e}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   657
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   658
definition
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   659
  cball :: "'a::metric_space \<Rightarrow> real \<Rightarrow> 'a set" where
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   660
  "cball x e = {y. dist x y \<le> e}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   661
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
   662
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
   663
  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
   664
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
   665
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
   666
  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
   667
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
   668
lemma mem_ball_0:
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   669
  fixes x :: "'a::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   670
  shows "x \<in> ball 0 e \<longleftrightarrow> norm x < e"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   671
  by (simp add: dist_norm)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   672
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
   673
lemma mem_cball_0:
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   674
  fixes x :: "'a::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   675
  shows "x \<in> cball 0 e \<longleftrightarrow> norm x \<le> e"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   676
  by (simp add: dist_norm)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   677
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
   678
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
   679
  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
   680
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 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
   682
  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
   683
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   684
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
   685
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
   686
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
   687
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
   688
  by (simp add: set_eq_iff) arith
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   689
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   690
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
   691
  by (simp add: set_eq_iff)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   692
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   693
lemma diff_less_iff: "(a::real) - b > 0 \<longleftrightarrow> a > b"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   694
  "(a::real) - b < 0 \<longleftrightarrow> a < b"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   695
  "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
   696
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
   697
  "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
   698
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   699
lemma open_ball[intro, simp]: "open (ball x e)"
44170
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   700
  unfolding open_dist ball_def mem_Collect_eq Ball_def
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   701
  unfolding dist_commute
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   702
  apply clarify
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   703
  apply (rule_tac x="e - dist xa x" in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   704
  using dist_triangle_alt[where z=x]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   705
  apply (clarsimp simp add: diff_less_iff)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   706
  apply atomize
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   707
  apply (erule_tac x="y" in allE)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   708
  apply (erule_tac x="xa" in allE)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   709
  by arith
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   710
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   711
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
   712
  unfolding open_dist subset_eq mem_ball Ball_def dist_commute ..
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   713
33714
eb2574ac4173 Added new lemmas to Euclidean Space by Robert Himmelmann
hoelzl
parents: 33324
diff changeset
   714
lemma openE[elim?]:
eb2574ac4173 Added new lemmas to Euclidean Space by Robert Himmelmann
hoelzl
parents: 33324
diff changeset
   715
  assumes "open S" "x\<in>S" 
eb2574ac4173 Added new lemmas to Euclidean Space by Robert Himmelmann
hoelzl
parents: 33324
diff changeset
   716
  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
   717
  using assms unfolding open_contains_ball by auto
eb2574ac4173 Added new lemmas to Euclidean Space by Robert Himmelmann
hoelzl
parents: 33324
diff changeset
   718
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   719
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
   720
  by (metis open_contains_ball subset_eq centre_in_ball)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   721
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   722
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
   723
  unfolding mem_ball set_eq_iff
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   724
  apply (simp add: not_less)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   725
  by (metis zero_le_dist order_trans dist_self)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   726
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   727
lemma ball_empty[intro]: "e \<le> 0 ==> ball x e = {}" by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   728
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
   729
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
   730
  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
   731
  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
   732
  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
   733
  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
   734
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
   735
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
   736
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   737
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
   738
  fixes x :: "'a\<Colon>euclidean_space"
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   739
  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
   740
  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
   741
proof -
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   742
  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
   743
  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
   744
  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
   745
  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
   746
    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
   747
  qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   748
  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
   749
  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
   750
  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
   751
    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
   752
  qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   753
  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
   754
  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
   755
  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
   756
  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
   757
    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
   758
    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
   759
      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
   760
    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
   761
    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
   762
      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
   763
      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
   764
      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
   765
      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
   766
      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
   767
      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
   768
        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
   769
      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
   770
        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
   771
      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
   772
        by (simp add: power_divide)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   773
    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
   774
    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
   775
    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
   776
  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
   777
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
   778
 
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
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
   780
  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
   781
  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
   782
  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
   783
  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
   784
  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
   785
  shows "M = (\<Union>f\<in>I. box (a' f) (b' f))"
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   786
proof safe
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   787
  fix x assume "x \<in> M"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   788
  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
   789
    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
   790
  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
   791
    "\<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
   792
    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
   793
  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
   794
     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
   795
        (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
   796
qed (auto simp: I_def)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   797
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   798
subsection{* Connectedness *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   799
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   800
definition "connected S \<longleftrightarrow>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   801
  ~(\<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
   802
  \<and> ~(e1 \<inter> S = {}) \<and> ~(e2 \<inter> S = {}))"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   803
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   804
lemma connected_local:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   805
 "connected S \<longleftrightarrow> ~(\<exists>e1 e2.
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   806
                 openin (subtopology euclidean S) e1 \<and>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   807
                 openin (subtopology euclidean S) e2 \<and>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   808
                 S \<subseteq> e1 \<union> e2 \<and>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   809
                 e1 \<inter> e2 = {} \<and>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   810
                 ~(e1 = {}) \<and>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   811
                 ~(e2 = {}))"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   812
unfolding connected_def openin_open by (safe, blast+)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   813
34105
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
   814
lemma exists_diff:
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
   815
  fixes P :: "'a set \<Rightarrow> bool"
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
   816
  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
   817
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   818
  {assume "?lhs" hence ?rhs by blast }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   819
  moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   820
  {fix S assume H: "P S"
34105
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
   821
    have "S = - (- S)" by auto
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
   822
    with H have "P (- (- S))" by metis }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   823
  ultimately show ?thesis by metis
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   824
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   825
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   826
lemma connected_clopen: "connected S \<longleftrightarrow>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   827
        (\<forall>T. openin (subtopology euclidean S) T \<and>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   828
            closedin (subtopology euclidean S) T \<longrightarrow> T = {} \<or> T = S)" (is "?lhs \<longleftrightarrow> ?rhs")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   829
proof-
34105
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
   830
  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
   831
    unfolding connected_def openin_open closedin_closed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   832
    apply (subst exists_diff) by blast
34105
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
   833
  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
   834
    (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
   835
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   836
  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
   837
    (is "_ \<longleftrightarrow> \<not> (\<exists>t' t. ?Q t' t)")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   838
    unfolding connected_def openin_open closedin_closed by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   839
  {fix e2
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   840
    {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
   841
        by auto}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   842
    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
   843
  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
   844
  then show ?thesis unfolding th0 th1 by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   845
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   846
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   847
lemma connected_empty[simp, intro]: "connected {}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   848
  by (simp add: connected_def)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   849
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   850
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   851
subsection{* Limit points *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   852
44207
ea99698c2070 locale-ize some definitions, so perfect_space and heine_borel can inherit from the proper superclasses
huffman
parents: 44170
diff changeset
   853
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
   854
  islimpt:: "'a \<Rightarrow> 'a set \<Rightarrow> bool" (infixr "islimpt" 60) where
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   855
  "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
   856
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   857
lemma islimptI:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   858
  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
   859
  shows "x islimpt S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   860
  using assms unfolding islimpt_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   861
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   862
lemma islimptE:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   863
  assumes "x islimpt S" and "x \<in> T" and "open T"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   864
  obtains y where "y \<in> S" and "y \<in> T" and "y \<noteq> x"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   865
  using assms unfolding islimpt_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   866
44584
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
   867
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
   868
  unfolding islimpt_def eventually_at_topological by auto
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
   869
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
   870
lemma islimpt_subset: "\<lbrakk>x islimpt S; S \<subseteq> T\<rbrakk> \<Longrightarrow> x islimpt T"
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
   871
  unfolding islimpt_def by fast
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   872
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   873
lemma islimpt_approachable:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   874
  fixes x :: "'a::metric_space"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   875
  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
   876
  unfolding islimpt_iff_eventually eventually_at by fast
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   877
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   878
lemma islimpt_approachable_le:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   879
  fixes x :: "'a::metric_space"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   880
  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
   881
  unfolding islimpt_approachable
44584
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
   882
  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
   883
    THEN arg_cong [where f=Not]]
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
   884
  by (simp add: Bex_def conj_commute conj_left_commute)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   885
44571
bd91b77c4cd6 move class perfect_space into RealVector.thy;
huffman
parents: 44568
diff changeset
   886
lemma islimpt_UNIV_iff: "x islimpt UNIV \<longleftrightarrow> \<not> open {x}"
bd91b77c4cd6 move class perfect_space into RealVector.thy;
huffman
parents: 44568
diff changeset
   887
  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
   888
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   889
text {* A perfect space has no isolated points. *}
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   890
44571
bd91b77c4cd6 move class perfect_space into RealVector.thy;
huffman
parents: 44568
diff changeset
   891
lemma islimpt_UNIV [simp, intro]: "(x::'a::perfect_space) islimpt UNIV"
bd91b77c4cd6 move class perfect_space into RealVector.thy;
huffman
parents: 44568
diff changeset
   892
  unfolding islimpt_UNIV_iff by (rule not_open_singleton)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   893
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   894
lemma perfect_choose_dist:
44072
5b970711fb39 class perfect_space inherits from topological_space;
huffman
parents: 43338
diff changeset
   895
  fixes x :: "'a::{perfect_space, metric_space}"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   896
  shows "0 < r \<Longrightarrow> \<exists>a. a \<noteq> x \<and> dist a x < r"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   897
using islimpt_UNIV [of x]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   898
by (simp add: islimpt_approachable)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   899
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   900
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
   901
  unfolding closed_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   902
  apply (subst open_subopen)
34105
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
   903
  apply (simp add: islimpt_def subset_eq)
44170
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   904
  by (metis ComplE ComplI)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   905
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   906
lemma islimpt_EMPTY[simp]: "\<not> x islimpt {}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   907
  unfolding islimpt_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   908
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   909
lemma finite_set_avoid:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   910
  fixes a :: "'a::metric_space"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   911
  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
   912
proof(induct rule: finite_induct[OF fS])
41863
e5104b436ea1 removed dependency on Dense_Linear_Order
boehmes
parents: 41413
diff changeset
   913
  case 1 thus ?case by (auto intro: zero_less_one)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   914
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   915
  case (2 x F)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   916
  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
   917
  {assume "x = a" hence ?case using d by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   918
  moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   919
  {assume xa: "x\<noteq>a"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   920
    let ?d = "min d (dist a x)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   921
    have dp: "?d > 0" using xa d(1) using dist_nz by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   922
    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
   923
    with dp xa have ?case by(auto intro!: exI[where x="?d"]) }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   924
  ultimately show ?case by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   925
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   926
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   927
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
   928
  by (simp add: islimpt_iff_eventually eventually_conj_iff)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   929
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   930
lemma discrete_imp_closed:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   931
  fixes S :: "'a::metric_space set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   932
  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
   933
  shows "closed S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   934
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   935
  {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
   936
    from e have e2: "e/2 > 0" by arith
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   937
    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
   938
    let ?m = "min (e/2) (dist x y) "
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   939
    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
   940
    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
   941
    have th: "dist z y < e" using z y
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   942
      by (intro dist_triangle_lt [where z=x], simp)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   943
    from d[rule_format, OF y(1) z(1) th] y z
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   944
    have False by (auto simp add: dist_commute)}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   945
  then show ?thesis by (metis islimpt_approachable closed_limpt [where 'a='a])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   946
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   947
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   948
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   949
subsection {* Interior of a Set *}
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   950
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   951
definition "interior S = \<Union>{T. open T \<and> T \<subseteq> S}"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   952
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   953
lemma interiorI [intro?]:
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   954
  assumes "open T" and "x \<in> T" and "T \<subseteq> S"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   955
  shows "x \<in> interior S"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   956
  using assms unfolding interior_def by fast
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   957
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   958
lemma interiorE [elim?]:
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   959
  assumes "x \<in> interior S"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   960
  obtains T where "open T" and "x \<in> T" and "T \<subseteq> S"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   961
  using assms unfolding interior_def by fast
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   962
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   963
lemma open_interior [simp, intro]: "open (interior S)"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   964
  by (simp add: interior_def open_Union)
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   965
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   966
lemma interior_subset: "interior S \<subseteq> S"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   967
  by (auto simp add: interior_def)
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   968
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   969
lemma interior_maximal: "T \<subseteq> S \<Longrightarrow> open T \<Longrightarrow> T \<subseteq> interior S"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   970
  by (auto simp add: interior_def)
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   971
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   972
lemma interior_open: "open S \<Longrightarrow> interior S = S"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   973
  by (intro equalityI interior_subset interior_maximal subset_refl)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   974
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   975
lemma interior_eq: "interior S = S \<longleftrightarrow> open S"
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   976
  by (metis open_interior interior_open)
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   977
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   978
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
   979
  by (metis interior_maximal interior_subset subset_trans)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   980
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   981
lemma interior_empty [simp]: "interior {} = {}"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   982
  using open_empty by (rule interior_open)
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   983
44522
2f7e9d890efe rename subset_{interior,closure} to {interior,closure}_mono;
huffman
parents: 44519
diff changeset
   984
lemma interior_UNIV [simp]: "interior UNIV = UNIV"
2f7e9d890efe rename subset_{interior,closure} to {interior,closure}_mono;
huffman
parents: 44519
diff changeset
   985
  using open_UNIV by (rule interior_open)
2f7e9d890efe rename subset_{interior,closure} to {interior,closure}_mono;
huffman
parents: 44519
diff changeset
   986
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   987
lemma interior_interior [simp]: "interior (interior S) = interior S"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   988
  using open_interior by (rule interior_open)
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   989
44522
2f7e9d890efe rename subset_{interior,closure} to {interior,closure}_mono;
huffman
parents: 44519
diff changeset
   990
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
   991
  by (auto simp add: interior_def)
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   992
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   993
lemma interior_unique:
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   994
  assumes "T \<subseteq> S" and "open T"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   995
  assumes "\<And>T'. T' \<subseteq> S \<Longrightarrow> open T' \<Longrightarrow> T' \<subseteq> T"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   996
  shows "interior S = T"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   997
  by (intro equalityI assms interior_subset open_interior interior_maximal)
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   998
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   999
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
  1000
  by (intro equalityI Int_mono Int_greatest interior_mono Int_lower1
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1001
    Int_lower2 interior_maximal interior_subset open_Int open_interior)
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1002
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1003
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
  1004
  using open_contains_ball_eq [where S="interior S"]
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1005
  by (simp add: open_subset_interior)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1006
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1007
lemma interior_limit_point [intro]:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1008
  fixes x :: "'a::perfect_space"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1009
  assumes x: "x \<in> interior S" shows "x islimpt S"
44072
5b970711fb39 class perfect_space inherits from topological_space;
huffman
parents: 43338
diff changeset
  1010
  using x islimpt_UNIV [of x]
5b970711fb39 class perfect_space inherits from topological_space;
huffman
parents: 43338
diff changeset
  1011
  unfolding interior_def islimpt_def
5b970711fb39 class perfect_space inherits from topological_space;
huffman
parents: 43338
diff changeset
  1012
  apply (clarsimp, rename_tac T T')
5b970711fb39 class perfect_space inherits from topological_space;
huffman
parents: 43338
diff changeset
  1013
  apply (drule_tac x="T \<inter> T'" in spec)
5b970711fb39 class perfect_space inherits from topological_space;
huffman
parents: 43338
diff changeset
  1014
  apply (auto simp add: open_Int)
5b970711fb39 class perfect_space inherits from topological_space;
huffman
parents: 43338
diff changeset
  1015
  done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1016
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1017
lemma interior_closed_Un_empty_interior:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1018
  assumes cS: "closed S" and iT: "interior T = {}"
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1019
  shows "interior (S \<union> T) = interior S"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1020
proof
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1021
  show "interior S \<subseteq> interior (S \<union> T)"
44522
2f7e9d890efe rename subset_{interior,closure} to {interior,closure}_mono;
huffman
parents: 44519
diff changeset
  1022
    by (rule interior_mono, rule Un_upper1)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1023
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1024
  show "interior (S \<union> T) \<subseteq> interior S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1025
  proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1026
    fix x assume "x \<in> interior (S \<union> T)"
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1027
    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
  1028
    show "x \<in> interior S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1029
    proof (rule ccontr)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1030
      assume "x \<notin> interior S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1031
      with `x \<in> R` `open R` obtain y where "y \<in> R - S"
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1032
        unfolding interior_def by fast
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1033
      from `open R` `closed S` have "open (R - S)" by (rule open_Diff)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1034
      from `R \<subseteq> S \<union> T` have "R - S \<subseteq> T" by fast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1035
      from `y \<in> R - S` `open (R - S)` `R - S \<subseteq> T` `interior T = {}`
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1036
      show "False" unfolding interior_def by fast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1037
    qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1038
  qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1039
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1040
44365
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1041
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
  1042
proof (rule interior_unique)
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1043
  show "interior A \<times> interior B \<subseteq> A \<times> B"
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1044
    by (intro Sigma_mono interior_subset)
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1045
  show "open (interior A \<times> interior B)"
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1046
    by (intro open_Times open_interior)
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1047
  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
  1048
  proof (safe)
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1049
    fix x y assume "(x, y) \<in> T"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1050
    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
  1051
      using `open T` unfolding open_prod_def by fast
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1052
    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
  1053
      using `T \<subseteq> A \<times> B` by auto
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1054
    thus "x \<in> interior A" and "y \<in> interior B"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1055
      by (auto intro: interiorI)
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1056
  qed
44365
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1057
qed
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1058
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1059
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1060
subsection {* Closure of a Set *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1061
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1062
definition "closure S = S \<union> {x | x. x islimpt S}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1063
44518
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
  1064
lemma interior_closure: "interior S = - (closure (- S))"
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
  1065
  unfolding interior_def closure_def islimpt_def by auto
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
  1066
34105
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
  1067
lemma closure_interior: "closure S = - interior (- S)"
44518
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
  1068
  unfolding interior_closure by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1069
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1070
lemma closed_closure[simp, intro]: "closed (closure S)"
44518
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
  1071
  unfolding closure_interior by (simp add: closed_Compl)
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
  1072
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
  1073
lemma closure_subset: "S \<subseteq> closure S"
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
  1074
  unfolding closure_def by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1075
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1076
lemma closure_hull: "closure S = closed hull S"
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1077
  unfolding hull_def closure_interior interior_def by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1078
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1079
lemma closure_eq: "closure S = S \<longleftrightarrow> closed S"
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1080
  unfolding closure_hull using closed_Inter by (rule hull_eq)
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1081
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1082
lemma closure_closed [simp]: "closed S \<Longrightarrow> closure S = S"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1083
  unfolding closure_eq .
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1084
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1085
lemma closure_closure [simp]: "closure (closure S) = closure S"
44518
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
  1086
  unfolding closure_hull by (rule hull_hull)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1087
44522
2f7e9d890efe rename subset_{interior,closure} to {interior,closure}_mono;
huffman
parents: 44519
diff changeset
  1088
lemma closure_mono: "S \<subseteq> T \<Longrightarrow> closure S \<subseteq> closure T"
44518
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
  1089
  unfolding closure_hull by (rule hull_mono)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1090
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1091
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
  1092
  unfolding closure_hull by (rule hull_minimal)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1093
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1094
lemma closure_unique:
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1095
  assumes "S \<subseteq> T" and "closed T"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1096
  assumes "\<And>T'. S \<subseteq> T' \<Longrightarrow> closed T' \<Longrightarrow> T \<subseteq> T'"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1097
  shows "closure S = T"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1098
  using assms unfolding closure_hull by (rule hull_unique)
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1099
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1100
lemma closure_empty [simp]: "closure {} = {}"
44518
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
  1101
  using closed_empty by (rule closure_closed)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1102
44522
2f7e9d890efe rename subset_{interior,closure} to {interior,closure}_mono;
huffman
parents: 44519
diff changeset
  1103
lemma closure_UNIV [simp]: "closure UNIV = UNIV"
44518
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
  1104
  using closed_UNIV by (rule closure_closed)
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
  1105
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
  1106
lemma closure_union [simp]: "closure (S \<union> T) = closure S \<union> closure T"
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
  1107
  unfolding closure_interior by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1108
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1109
lemma closure_eq_empty: "closure S = {} \<longleftrightarrow> S = {}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1110
  using closure_empty closure_subset[of S]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1111
  by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1112
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1113
lemma closure_subset_eq: "closure S \<subseteq> S \<longleftrightarrow> closed S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1114
  using closure_eq[of S] closure_subset[of S]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1115
  by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1116
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1117
lemma open_inter_closure_eq_empty:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1118
  "open S \<Longrightarrow> (S \<inter> closure T) = {} \<longleftrightarrow> S \<inter> T = {}"
34105
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
  1119
  using open_subset_interior[of S "- T"]
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
  1120
  using interior_subset[of "- T"]
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1121
  unfolding closure_interior
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1122
  by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1123
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1124
lemma open_inter_closure_subset:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1125
  "open S \<Longrightarrow> (S \<inter> (closure T)) \<subseteq> closure(S \<inter> T)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1126
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1127
  fix x
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1128
  assume as: "open S" "x \<in> S \<inter> closure T"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1129
  { assume *:"x islimpt T"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1130
    have "x islimpt (S \<inter> T)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1131
    proof (rule islimptI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1132
      fix A
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1133
      assume "x \<in> A" "open A"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1134
      with as have "x \<in> A \<inter> S" "open (A \<inter> S)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1135
        by (simp_all add: open_Int)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1136
      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
  1137
        by (rule islimptE)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1138
      hence "y \<in> S \<inter> T" "y \<in> A \<and> y \<noteq> x"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1139
        by simp_all
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1140
      thus "\<exists>y\<in>(S \<inter> T). y \<in> A \<and> y \<noteq> x" ..
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1141
    qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1142
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1143
  then show "x \<in> closure (S \<inter> T)" using as
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1144
    unfolding closure_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1145
    by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1146
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1147
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1148
lemma closure_complement: "closure (- S) = - interior S"
44518
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
  1149
  unfolding closure_interior by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1150
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1151
lemma interior_complement: "interior (- S) = - closure S"
44518
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
  1152
  unfolding closure_interior by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1153
44365
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1154
lemma closure_Times: "closure (A \<times> B) = closure A \<times> closure B"
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1155
proof (rule closure_unique)
44365
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1156
  show "A \<times> B \<subseteq> closure A \<times> closure B"
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1157
    by (intro Sigma_mono closure_subset)
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1158
  show "closed (closure A \<times> closure B)"
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1159
    by (intro closed_Times closed_closure)
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1160
  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
  1161
    apply (simp add: closed_def open_prod_def, clarify)
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1162
    apply (rule ccontr)
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1163
    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
  1164
    apply (simp add: closure_interior interior_def)
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1165
    apply (drule_tac x=C in spec)
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1166
    apply (drule_tac x=D in spec)
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1167
    apply auto
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1168
    done
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1169
qed
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1170
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1171
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1172
subsection {* Frontier (aka boundary) *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1173
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1174
definition "frontier S = closure S - interior S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1176
lemma frontier_closed: "closed(frontier S)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1177
  by (simp add: frontier_def closed_Diff)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1178
34105
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
  1179
lemma frontier_closures: "frontier S = (closure S) \<inter> (closure(- S))"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1180
  by (auto simp add: frontier_def interior_closure)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1181
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1182
lemma frontier_straddle:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1183
  fixes a :: "'a::metric_space"
44909
1f5d6eb73549 shorten proof of frontier_straddle
huffman
parents: 44907
diff changeset
  1184
  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
  1185
  unfolding frontier_def closure_interior
1f5d6eb73549 shorten proof of frontier_straddle
huffman
parents: 44907
diff changeset
  1186
  by (auto simp add: mem_interior subset_eq ball_def)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1187
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1188
lemma frontier_subset_closed: "closed S \<Longrightarrow> frontier S \<subseteq> S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1189
  by (metis frontier_def closure_closed Diff_subset)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1190
34964
4e8be3c04d37 Replaced vec1 and dest_vec1 by abbreviation.
hoelzl
parents: 34291
diff changeset
  1191
lemma frontier_empty[simp]: "frontier {} = {}"
36362
06475a1547cb fix lots of looping simp calls and other warnings
huffman
parents: 36360
diff changeset
  1192
  by (simp add: frontier_def)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1193
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1194
lemma frontier_subset_eq: "frontier S \<subseteq> S \<longleftrightarrow> closed S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1195
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1196
  { assume "frontier S \<subseteq> S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1197
    hence "closure S \<subseteq> S" using interior_subset unfolding frontier_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1198
    hence "closed S" using closure_subset_eq by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1199
  }
36362
06475a1547cb fix lots of looping simp calls and other warnings
huffman
parents: 36360
diff changeset
  1200
  thus ?thesis using frontier_subset_closed[of S] ..
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1201
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1202
34105
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
  1203
lemma frontier_complement: "frontier(- S) = frontier S"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1204
  by (auto simp add: frontier_def closure_complement interior_complement)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1205
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1206
lemma frontier_disjoint_eq: "frontier S \<inter> S = {} \<longleftrightarrow> open S"
34105
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
  1207
  using frontier_complement frontier_subset_eq[of "- S"]
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
  1208
  unfolding open_closed by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1209
44081
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44076
diff changeset
  1210
subsection {* Filters and the ``eventually true'' quantifier *}
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44076
diff changeset
  1211
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1212
definition
44081
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44076
diff changeset
  1213
  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
  1214
    (infixr "indirection" 70) where
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1215
  "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
  1216
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  1217
text {* Identify Trivial limits, where we can't approach arbitrarily closely. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1218
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1219
lemma trivial_limit_within:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1220
  shows "trivial_limit (at a within S) \<longleftrightarrow> \<not> a islimpt S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1221
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1222
  assume "trivial_limit (at a within S)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1223
  thus "\<not> a islimpt S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1224
    unfolding trivial_limit_def
36358
246493d61204 define nets directly as filters, instead of as filter bases
huffman
parents: 36336
diff changeset
  1225
    unfolding eventually_within eventually_at_topological
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1226
    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
  1227
    apply (clarsimp simp add: set_eq_iff)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1228
    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
  1229
    apply (clarsimp, drule_tac x=y in bspec, simp_all)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1230
    done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1231
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1232
  assume "\<not> a islimpt S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1233
  thus "trivial_limit (at a within S)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1234
    unfolding trivial_limit_def
36358
246493d61204 define nets directly as filters, instead of as filter bases
huffman
parents: 36336
diff changeset
  1235
    unfolding eventually_within eventually_at_topological
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1236
    unfolding islimpt_def
36358
246493d61204 define nets directly as filters, instead of as filter bases
huffman
parents: 36336
diff changeset
  1237
    apply clarsimp
246493d61204 define nets directly as filters, instead of as filter bases
huffman
parents: 36336
diff changeset
  1238
    apply (rule_tac x=T in exI)
246493d61204 define nets directly as filters, instead of as filter bases
huffman
parents: 36336
diff changeset
  1239
    apply auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1240
    done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1241
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1242
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1243
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
  1244
  using trivial_limit_within [of a UNIV] by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1245
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1246
lemma trivial_limit_at:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1247
  fixes a :: "'a::perfect_space"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1248
  shows "\<not> trivial_limit (at a)"
44571
bd91b77c4cd6 move class perfect_space into RealVector.thy;
huffman
parents: 44568
diff changeset
  1249
  by (rule at_neq_bot)
33175
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_infinity:
44081
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44076
diff changeset
  1252
  "\<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
  1253
  unfolding trivial_limit_def eventually_at_infinity
246493d61204 define nets directly as filters, instead of as filter bases
huffman
parents: 36336
diff changeset
  1254
  apply clarsimp
44072
5b970711fb39 class perfect_space inherits from topological_space;
huffman
parents: 43338
diff changeset
  1255
  apply (subgoal_tac "\<exists>x::'a. x \<noteq> 0", clarify)
5b970711fb39 class perfect_space inherits from topological_space;
huffman
parents: 43338
diff changeset
  1256
   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
  1257
  apply (cut_tac islimpt_UNIV [of "0::'a", unfolded islimpt_def])
5b970711fb39 class perfect_space inherits from topological_space;
huffman
parents: 43338
diff changeset
  1258
  apply (drule_tac x=UNIV in spec, simp)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1259
  done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1260
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  1261
text {* Some property holds "sufficiently close" to the limit point. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1262
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1263
lemma eventually_at: (* FIXME: this replaces Limits.eventually_at *)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1264
  "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
  1265
unfolding eventually_at dist_nz by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1266
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
  1267
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
  1268
  "eventually P (at a within S) \<longleftrightarrow>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1269
        (\<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
  1270
  by (rule eventually_within_less)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1271
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1272
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
  1273
  unfolding trivial_limit_def
246493d61204 define nets directly as filters, instead of as filter bases
huffman
parents: 36336
diff changeset
  1274
  by (auto elim: eventually_rev_mp)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1275
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1276
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
  1277
  by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1278
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1279
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
  1280
  by (simp add: filter_eq_iff)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1281
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1282
text{* Combining theorems for "eventually" *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1283
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1284
lemma eventually_rev_mono:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1285
  "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
  1286
using eventually_mono [of P Q] by fast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1287
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1288
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
  1289
  by (simp add: eventually_False)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1290
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1291
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  1292
subsection {* Limits *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1293
44081
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44076
diff changeset
  1294
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
  1295
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44076
diff changeset
  1296
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
  1297
  where "Lim A f = (THE l. (f ---> l) A)"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1298
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1299
lemma Lim:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1300
 "(f ---> l) net \<longleftrightarrow>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1301
        trivial_limit net \<or>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1302
        (\<forall>e>0. eventually (\<lambda>x. dist (f x) l < e) net)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1303
  unfolding tendsto_iff trivial_limit_eq by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1304
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1305
text{* Show that they yield usual definitions in the various cases. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1306
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1307
lemma Lim_within_le: "(f ---> l)(at a within S) \<longleftrightarrow>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1308
           (\<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
  1309
  by (auto simp add: tendsto_iff eventually_within_le)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1310
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1311
lemma Lim_within: "(f ---> l) (at a within S) \<longleftrightarrow>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1312
        (\<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
  1313
  by (auto simp add: tendsto_iff eventually_within)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1314
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1315
lemma Lim_at: "(f ---> l) (at a) \<longleftrightarrow>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1316
        (\<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
  1317
  by (auto simp add: tendsto_iff eventually_at)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1318
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1319
lemma Lim_at_infinity:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1320
  "(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
  1321
  by (auto simp add: tendsto_iff eventually_at_infinity)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1322
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1323
lemma Lim_eventually: "eventually (\<lambda>x. f x = l) net \<Longrightarrow> (f ---> l) net"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1324
  by (rule topological_tendstoI, auto elim: eventually_rev_mono)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1325
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1326
text{* The expected monotonicity property. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1327
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1328
lemma Lim_within_empty: "(f ---> l) (net within {})"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1329
  unfolding tendsto_def Limits.eventually_within by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1330
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1331
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
  1332
  unfolding tendsto_def Limits.eventually_within
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1333
  by (auto elim!: eventually_elim1)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1334
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1335
lemma Lim_Un: assumes "(f ---> l) (net within S)" "(f ---> l) (net within T)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1336
  shows "(f ---> l) (net within (S \<union> T))"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1337
  using assms unfolding tendsto_def Limits.eventually_within
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1338
  apply clarify
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1339
  apply (drule spec, drule (1) mp, drule (1) mp)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1340
  apply (drule spec, drule (1) mp, drule (1) mp)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1341
  apply (auto elim: eventually_elim2)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1342
  done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1343
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1344
lemma Lim_Un_univ:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1345
 "(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
  1346
        ==> (f ---> l) net"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1347
  by (metis Lim_Un within_UNIV)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1348
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1349
text{* Interrelations between restricted and unrestricted limits. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1350
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1351
lemma Lim_at_within: "(f ---> l) net ==> (f ---> l)(net within S)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1352
  (* FIXME: rename *)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1353
  unfolding tendsto_def Limits.eventually_within
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1354
  apply (clarify, drule spec, drule (1) mp, drule (1) mp)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1355
  by (auto elim!: eventually_elim1)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1356
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1357
lemma eventually_within_interior:
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1358
  assumes "x \<in> interior S"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1359
  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
  1360
proof-
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1361
  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
  1362
  { assume "?lhs"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1363
    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
  1364
      unfolding Limits.eventually_within Limits.eventually_at_topological
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1365
      by auto
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1366
    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
  1367
      by auto
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1368
    then have "?rhs"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1369
      unfolding Limits.eventually_at_topological by auto
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1370
  } moreover
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1371
  { assume "?rhs" hence "?lhs"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1372
      unfolding Limits.eventually_within
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1373
      by (auto elim: eventually_elim1)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1374
  } ultimately
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1375
  show "?thesis" ..
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1376
qed
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1377
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1378
lemma at_within_interior:
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1379
  "x \<in> interior S \<Longrightarrow> at x within S = at x"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1380
  by (simp add: filter_eq_iff eventually_within_interior)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1381
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1382
lemma at_within_open:
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1383
  "\<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
  1384
  by (simp only: at_within_interior interior_open)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1385
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1386
lemma Lim_within_open:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1387
  fixes f :: "'a::topological_space \<Rightarrow> 'b::topological_space"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1388
  assumes"a \<in> S" "open S"
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1389
  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
  1390
  using assms by (simp only: at_within_open)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1391
43338
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1392
lemma Lim_within_LIMSEQ:
44584
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  1393
  fixes a :: "'a::metric_space"
43338
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1394
  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
  1395
  shows "(X ---> L) (at a within T)"
44584
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  1396
  using assms unfolding tendsto_def [where l=L]
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  1397
  by (simp add: sequentially_imp_eventually_within)
43338
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1398
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1399
lemma Lim_right_bound:
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1400
  fixes f :: "real \<Rightarrow> real"
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1401
  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
  1402
  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
  1403
  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
  1404
proof cases
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1405
  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
  1406
next
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1407
  assume [simp]: "{x<..} \<inter> I \<noteq> {}"
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1408
  show ?thesis
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1409
  proof (rule Lim_within_LIMSEQ, safe)
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1410
    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
  1411
    
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1412
    show "(\<lambda>n. f (S n)) ----> Inf (f ` ({x<..} \<inter> I))"
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1413
    proof (rule LIMSEQ_I, rule ccontr)
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1414
      fix r :: real assume "0 < r"
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1415
      with Inf_close[of "f ` ({x<..} \<inter> I)" r]
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1416
      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
  1417
      from `x < y` have "0 < y - x" by auto
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1418
      from S(2)[THEN LIMSEQ_D, OF this]
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1419
      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
  1420
      
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1421
      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
  1422
      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
  1423
        using S bnd by (intro Inf_lower[where z=K]) auto
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1424
      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
  1425
        by (auto simp: not_less field_simps)
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1426
      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
  1427
      show False by auto
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1428
    qed
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1429
  qed
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1430
qed
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1431
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1432
text{* Another limit point characterization. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1433
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1434
lemma islimpt_sequential:
50883
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1435
  fixes x :: "'a::first_countable_topology"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1436
  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
  1437
    (is "?lhs = ?rhs")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1438
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1439
  assume ?lhs
50883
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1440
  from countable_basis_at_decseq[of x] guess A . note A = this
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1441
  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
  1442
  { fix n
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1443
    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
  1444
      unfolding islimpt_def using A(1,2)[of n] by auto
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1445
    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
  1446
      unfolding f_def by (rule someI_ex)
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1447
    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
  1448
  then have "\<forall>n. f n \<in> S - {x}" by auto
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1449
  moreover have "(\<lambda>n. f n) ----> x"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1450
  proof (rule topological_tendstoI)
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1451
    fix S assume "open S" "x \<in> S"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1452
    from A(3)[OF this] `\<And>n. f n \<in> A n`
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1453
    show "eventually (\<lambda>x. f x \<in> S) sequentially" by (auto elim!: eventually_elim1)
44584
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  1454
  qed
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  1455
  ultimately show ?rhs by fast
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1456
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1457
  assume ?rhs
50883
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1458
  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
  1459
  show ?lhs
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1460
    unfolding islimpt_def
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1461
  proof safe
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1462
    fix T assume "open T" "x \<in> T"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1463
    from lim[THEN topological_tendstoD, OF this] f
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1464
    show "\<exists>y\<in>S. y \<in> T \<and> y \<noteq> x"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1465
      unfolding eventually_sequentially by auto
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1466
  qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1467
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1468
44125
230a8665c919 mark some redundant theorems as legacy
huffman
parents: 44122
diff changeset
  1469
lemma Lim_inv: (* TODO: delete *)
44081
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44076
diff changeset
  1470
  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
  1471
  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
  1472
  shows "((inverse o f) ---> inverse l) A"
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  1473
  unfolding o_def using assms by (rule tendsto_inverse)
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  1474
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1475
lemma Lim_null:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1476
  fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
44125
230a8665c919 mark some redundant theorems as legacy
huffman
parents: 44122
diff changeset
  1477
  shows "(f ---> l) net \<longleftrightarrow> ((\<lambda>x. f(x) - l) ---> 0) net"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1478
  by (simp add: Lim dist_norm)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1479
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1480
lemma Lim_null_comparison:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1481
  fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1482
  assumes "eventually (\<lambda>x. norm (f x) \<le> g x) net" "(g ---> 0) net"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1483
  shows "(f ---> 0) net"
44252
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1484
proof (rule metric_tendsto_imp_tendsto)
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1485
  show "(g ---> 0) net" by fact
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1486
  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
  1487
    using assms(1) by (rule eventually_elim1, simp add: dist_norm)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1488
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1489
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1490
lemma Lim_transform_bound:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1491
  fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1492
  fixes g :: "'a \<Rightarrow> 'c::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1493
  assumes "eventually (\<lambda>n. norm(f n) <= norm(g n)) net"  "(g ---> 0) net"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1494
  shows "(f ---> 0) net"
44252
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1495
  using assms(1) tendsto_norm_zero [OF assms(2)]
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1496
  by (rule Lim_null_comparison)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1497
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1498
text{* Deducing things about the limit from the elements. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1499
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1500
lemma Lim_in_closed_set:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1501
  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
  1502
  shows "l \<in> S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1503
proof (rule ccontr)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1504
  assume "l \<notin> S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1505
  with `closed S` have "open (- S)" "l \<in> - S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1506
    by (simp_all add: open_Compl)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1507
  with assms(4) have "eventually (\<lambda>x. f x \<in> - S) net"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1508
    by (rule topological_tendstoD)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1509
  with assms(2) have "eventually (\<lambda>x. False) net"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1510
    by (rule eventually_elim2) simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1511
  with assms(3) show "False"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1512
    by (simp add: eventually_False)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1513
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1514
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1515
text{* Need to prove closed(cball(x,e)) before deducing this as a corollary. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1516
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1517
lemma Lim_dist_ubound:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1518
  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
  1519
  shows "dist a l <= e"
44252
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1520
proof-
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1521
  have "dist a l \<in> {..e}"
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1522
  proof (rule Lim_in_closed_set)
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1523
    show "closed {..e}" by simp
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1524
    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
  1525
    show "\<not> trivial_limit net" by fact
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1526
    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
  1527
  qed
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1528
  thus ?thesis by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1529
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1530
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1531
lemma Lim_norm_ubound:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1532
  fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1533
  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
  1534
  shows "norm(l) <= e"
44252
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1535
proof-
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1536
  have "norm l \<in> {..e}"
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1537
  proof (rule Lim_in_closed_set)
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1538
    show "closed {..e}" by simp
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1539
    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
  1540
    show "\<not> trivial_limit net" by fact
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1541
    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
  1542
  qed
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1543
  thus ?thesis by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1544
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1545
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1546
lemma Lim_norm_lbound:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1547
  fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1548
  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
  1549
  shows "e \<le> norm l"
44252
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1550
proof-
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1551
  have "norm l \<in> {e..}"
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1552
  proof (rule Lim_in_closed_set)
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1553
    show "closed {e..}" by simp
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1554
    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
  1555
    show "\<not> trivial_limit net" by fact
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1556
    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
  1557
  qed
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1558
  thus ?thesis by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1559
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1560
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1561
text{* Uniqueness of the limit, when nontrivial. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1562
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1563
lemma tendsto_Lim:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1564
  fixes f :: "'a \<Rightarrow> 'b::t2_space"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1565
  shows "~(trivial_limit net) \<Longrightarrow> (f ---> l) net ==> Lim net f = l"
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41969
diff changeset
  1566
  unfolding Lim_def using tendsto_unique[of net f] by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1567
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1568
text{* Limit under bilinear function *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1569
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1570
lemma Lim_bilinear:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1571
  assumes "(f ---> l) net" and "(g ---> m) net" and "bounded_bilinear h"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1572
  shows "((\<lambda>x. h (f x) (g x)) ---> (h l m)) net"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1573
using `bounded_bilinear h` `(f ---> l) net` `(g ---> m) net`
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1574
by (rule bounded_bilinear.tendsto)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1575
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1576
text{* These are special for limits out of the same vector space. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1577
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1578
lemma Lim_within_id: "(id ---> a) (at a within s)"
45031
9583f2b56f85 add lemmas within_empty and tendsto_bot;
huffman
parents: 44909
diff changeset
  1579
  unfolding id_def by (rule tendsto_ident_at_within)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1580
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1581
lemma Lim_at_id: "(id ---> a) (at a)"
45031
9583f2b56f85 add lemmas within_empty and tendsto_bot;
huffman
parents: 44909
diff changeset
  1582
  unfolding id_def by (rule tendsto_ident_at)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1583
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1584
lemma Lim_at_zero:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1585
  fixes a :: "'a::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1586
  fixes l :: "'b::topological_space"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1587
  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
  1588
  using LIM_offset_zero LIM_offset_zero_cancel ..
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1589
44081
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44076
diff changeset
  1590
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
  1591
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1592
definition
44081
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44076
diff changeset
  1593
  netlimit :: "'a::t2_space filter \<Rightarrow> 'a" where
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1594
  "netlimit net = (SOME a. ((\<lambda>x. x) ---> a) net)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1595
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1596
lemma netlimit_within:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1597
  assumes "\<not> trivial_limit (at a within S)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1598
  shows "netlimit (at a within S) = a"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1599
unfolding netlimit_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1600
apply (rule some_equality)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1601
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
  1602
apply (rule tendsto_ident_at)
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41969
diff changeset
  1603
apply (erule tendsto_unique [OF assms])
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1604
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
  1605
apply (rule tendsto_ident_at)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1606
done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1607
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1608
lemma netlimit_at:
44072
5b970711fb39 class perfect_space inherits from topological_space;
huffman
parents: 43338
diff changeset
  1609
  fixes a :: "'a::{perfect_space,t2_space}"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1610
  shows "netlimit (at a) = a"
45031
9583f2b56f85 add lemmas within_empty and tendsto_bot;
huffman
parents: 44909
diff changeset
  1611
  using netlimit_within [of a UNIV] by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1612
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1613
lemma lim_within_interior:
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1614
  "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
  1615
  by (simp add: at_within_interior)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1616
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1617
lemma netlimit_within_interior:
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1618
  fixes x :: "'a::{t2_space,perfect_space}"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1619
  assumes "x \<in> interior S"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1620
  shows "netlimit (at x within S) = x"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1621
using assms by (simp add: at_within_interior netlimit_at)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1622
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1623
text{* Transformation of limit. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1624
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1625
lemma Lim_transform:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1626
  fixes f g :: "'a::type \<Rightarrow> 'b::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1627
  assumes "((\<lambda>x. f x - g x) ---> 0) net" "(f ---> l) net"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1628
  shows "(g ---> l) net"
44252
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1629
  using tendsto_diff [OF assms(2) assms(1)] by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1630
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1631
lemma Lim_transform_eventually:
36667
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1632
  "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
  1633
  apply (rule topological_tendstoI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1634
  apply (drule (2) topological_tendstoD)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1635
  apply (erule (1) eventually_elim2, simp)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1636
  done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1637
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1638
lemma Lim_transform_within:
36667
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1639
  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
  1640
  and "(f ---> l) (at x within S)"
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1641
  shows "(g ---> l) (at x within S)"
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1642
proof (rule Lim_transform_eventually)
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1643
  show "eventually (\<lambda>x. f x = g x) (at x within S)"
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1644
    unfolding eventually_within
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1645
    using assms(1,2) by auto
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1646
  show "(f ---> l) (at x within S)" by fact
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1647
qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1648
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1649
lemma Lim_transform_at:
36667
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1650
  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
  1651
  and "(f ---> l) (at x)"
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1652
  shows "(g ---> l) (at x)"
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1653
proof (rule Lim_transform_eventually)
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1654
  show "eventually (\<lambda>x. f x = g x) (at x)"
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1655
    unfolding eventually_at
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1656
    using assms(1,2) by auto
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1657
  show "(f ---> l) (at x)" by fact
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1658
qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1659
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1660
text{* Common case assuming being away from some crucial point like 0. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1661
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1662
lemma Lim_transform_away_within:
36669
c90c8a3ae1f7 generalize some lemmas to class t1_space
huffman
parents: 36668
diff changeset
  1663
  fixes a b :: "'a::t1_space"
36667
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1664
  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
  1665
  and "(f ---> l) (at a within S)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1666
  shows "(g ---> l) (at a within S)"
36669
c90c8a3ae1f7 generalize some lemmas to class t1_space
huffman
parents: 36668
diff changeset
  1667
proof (rule Lim_transform_eventually)
c90c8a3ae1f7 generalize some lemmas to class t1_space
huffman
parents: 36668
diff changeset
  1668
  show "(f ---> l) (at a within S)" by fact
c90c8a3ae1f7 generalize some lemmas to class t1_space
huffman
parents: 36668
diff changeset
  1669
  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
  1670
    unfolding Limits.eventually_within eventually_at_topological
c90c8a3ae1f7 generalize some lemmas to class t1_space
huffman
parents: 36668
diff changeset
  1671
    by (rule exI [where x="- {b}"], simp add: open_Compl assms)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1672
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1673
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1674
lemma Lim_transform_away_at:
36669
c90c8a3ae1f7 generalize some lemmas to class t1_space
huffman
parents: 36668
diff changeset
  1675
  fixes a b :: "'a::t1_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1676
  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
  1677
  and fl: "(f ---> l) (at a)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1678
  shows "(g ---> l) (at a)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1679
  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
  1680
  by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1681
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1682
text{* Alternatively, within an open set. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1683
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1684
lemma Lim_transform_within_open:
36667
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1685
  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
  1686
  and "(f ---> l) (at a)"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1687
  shows "(g ---> l) (at a)"
36667
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1688
proof (rule Lim_transform_eventually)
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1689
  show "eventually (\<lambda>x. f x = g x) (at a)"
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1690
    unfolding eventually_at_topological
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1691
    using assms(1,2,3) by auto
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1692
  show "(f ---> l) (at a)" by fact
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1693
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1694
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1695
text{* A congruence rule allowing us to transform limits assuming not at point. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1696
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1697
(* FIXME: Only one congruence rule for tendsto can be used at a time! *)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1698
36362
06475a1547cb fix lots of looping simp calls and other warnings
huffman
parents: 36360
diff changeset
  1699
lemma Lim_cong_within(*[cong add]*):
43338
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1700
  assumes "a = b" "x = y" "S = T"
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1701
  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
  1702
  shows "(f ---> x) (at a within S) \<longleftrightarrow> (g ---> y) (at b within T)"
36667
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1703
  unfolding tendsto_def Limits.eventually_within eventually_at_topological
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1704
  using assms by simp
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1705
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1706
lemma Lim_cong_at(*[cong add]*):
43338
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1707
  assumes "a = b" "x = y"
36667
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1708
  assumes "\<And>x. x \<noteq> a \<Longrightarrow> f x = g x"
43338
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1709
  shows "((\<lambda>x. f x) ---> x) (at a) \<longleftrightarrow> ((g ---> y) (at a))"
36667
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1710
  unfolding tendsto_def eventually_at_topological
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1711
  using assms by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1712
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1713
text{* Useful lemmas on closure and set of possible sequential limits.*}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1714
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1715
lemma closure_sequential:
50883
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1716
  fixes l :: "'a::first_countable_topology"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1717
  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
  1718
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1719
  assume "?lhs" moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1720
  { assume "l \<in> S"
44125
230a8665c919 mark some redundant theorems as legacy
huffman
parents: 44122
diff changeset
  1721
    hence "?rhs" using tendsto_const[of l sequentially] by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1722
  } moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1723
  { assume "l islimpt S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1724
    hence "?rhs" unfolding islimpt_sequential by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1725
  } ultimately
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1726
  show "?rhs" unfolding closure_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1727
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1728
  assume "?rhs"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1729
  thus "?lhs" unfolding closure_def unfolding islimpt_sequential by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1730
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1731
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1732
lemma closed_sequential_limits:
50883
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1733
  fixes S :: "'a::first_countable_topology set"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1734
  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
  1735
  unfolding closed_limpt
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1736
  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
  1737
  by metis
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1738
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1739
lemma closure_approachable:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1740
  fixes S :: "'a::metric_space set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1741
  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
  1742
  apply (auto simp add: closure_def islimpt_approachable)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1743
  by (metis dist_self)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1744
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1745
lemma closed_approachable:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1746
  fixes S :: "'a::metric_space set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1747
  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
  1748
  by (metis closure_closed closure_approachable)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1749
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1750
subsection {* Infimum Distance *}
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1751
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1752
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
  1753
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1754
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
  1755
  by (simp add: infdist_def)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1756
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1757
lemma infdist_nonneg:
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1758
  shows "0 \<le> infdist x A"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1759
  using assms by (auto simp add: infdist_def)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1760
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1761
lemma infdist_le:
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1762
  assumes "a \<in> A"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1763
  assumes "d = dist x a"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1764
  shows "infdist x A \<le> d"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1765
  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
  1766
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1767
lemma infdist_zero[simp]:
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1768
  assumes "a \<in> A" shows "infdist a A = 0"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1769
proof -
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1770
  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
  1771
  with infdist_nonneg[of a A] assms show "infdist a A = 0" by auto
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1772
qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1773
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1774
lemma infdist_triangle:
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1775
  shows "infdist x A \<le> infdist y A + dist x y"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1776
proof cases
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1777
  assume "A = {}" thus ?thesis by (simp add: infdist_def)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1778
next
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1779
  assume "A \<noteq> {}" then obtain a where "a \<in> A" by auto
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1780
  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
  1781
  proof
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1782
    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
  1783
    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
  1784
    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
  1785
    show "infdist x A \<le> d"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1786
      unfolding infdist_notempty[OF `A \<noteq> {}`]
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1787
    proof (rule Inf_lower2)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1788
      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
  1789
      show "dist x a \<le> d" unfolding d by (rule dist_triangle)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1790
      fix d assume "d \<in> {dist x a |a. a \<in> A}"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1791
      then obtain a where "a \<in> A" "d = dist x a" by auto
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1792
      thus "infdist x A \<le> d" by (rule infdist_le)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1793
    qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1794
  qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1795
  also have "\<dots> = dist x y + infdist y A"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1796
  proof (rule Inf_eq, safe)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1797
    fix a assume "a \<in> A"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1798
    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
  1799
  next
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1800
    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
  1801
    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
  1802
      by (intro Inf_greatest) (auto simp: field_simps)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1803
    thus "i \<le> dist x y + infdist y A" by simp
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1804
  qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1805
  finally show ?thesis by simp
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1806
qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1807
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1808
lemma
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1809
  in_closure_iff_infdist_zero:
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1810
  assumes "A \<noteq> {}"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1811
  shows "x \<in> closure A \<longleftrightarrow> infdist x A = 0"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1812
proof
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1813
  assume "x \<in> closure A"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1814
  show "infdist x A = 0"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1815
  proof (rule ccontr)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1816
    assume "infdist x A \<noteq> 0"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1817
    with infdist_nonneg[of x A] have "infdist x A > 0" by auto
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1818
    hence "ball x (infdist x A) \<inter> closure A = {}" apply auto
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1819
      by (metis `0 < infdist x A` `x \<in> closure A` closure_approachable dist_commute
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1820
        eucl_less_not_refl euclidean_trans(2) infdist_le)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1821
    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
  1822
    thus False using `x \<in> closure A` by simp
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1823
  qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1824
next
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1825
  assume x: "infdist x A = 0"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1826
  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
  1827
  show "x \<in> closure A" unfolding closure_approachable
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1828
  proof (safe, rule ccontr)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1829
    fix e::real assume "0 < e"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1830
    assume "\<not> (\<exists>y\<in>A. dist y x < e)"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1831
    hence "infdist x A \<ge> e" using `a \<in> A`
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1832
      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
  1833
      by (force simp: dist_commute)
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1834
    with x `0 < e` show False by auto
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1835
  qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1836
qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1837
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1838
lemma
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1839
  in_closed_iff_infdist_zero:
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1840
  assumes "closed A" "A \<noteq> {}"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1841
  shows "x \<in> A \<longleftrightarrow> infdist x A = 0"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1842
proof -
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1843
  have "x \<in> closure A \<longleftrightarrow> infdist x A = 0"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1844
    by (rule in_closure_iff_infdist_zero) fact
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1845
  with assms show ?thesis by simp
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1846
qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1847
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1848
lemma tendsto_infdist [tendsto_intros]:
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1849
  assumes f: "(f ---> l) F"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1850
  shows "((\<lambda>x. infdist (f x) A) ---> infdist l A) F"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1851
proof (rule tendstoI)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1852
  fix e ::real assume "0 < e"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1853
  from tendstoD[OF f this]
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1854
  show "eventually (\<lambda>x. dist (infdist (f x) A) (infdist l A) < e) F"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1855
  proof (eventually_elim)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1856
    fix x
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1857
    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
  1858
    have "dist (infdist (f x) A) (infdist l A) \<le> dist (f x) l"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1859
      by (simp add: dist_commute dist_real_def)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1860
    also assume "dist (f x) l < e"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1861
    finally show "dist (infdist (f x) A) (infdist l A) < e" .
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1862
  qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1863
qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1864
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1865
text{* Some other lemmas about sequences. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1866
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  1867
lemma sequentially_offset:
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  1868
  assumes "eventually (\<lambda>i. P i) sequentially"
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  1869
  shows "eventually (\<lambda>i. P (i + k)) sequentially"
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  1870
  using assms unfolding eventually_sequentially by (metis trans_le_add1)
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  1871
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1872
lemma seq_offset:
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  1873
  assumes "(f ---> l) sequentially"
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  1874
  shows "((\<lambda>i. f (i + k)) ---> l) sequentially"
44584
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  1875
  using assms by (rule LIMSEQ_ignore_initial_segment) (* FIXME: redundant *)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1876
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1877
lemma seq_offset_neg:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1878
  "(f ---> l) sequentially ==> ((\<lambda>i. f(i - k)) ---> l) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1879
  apply (rule topological_tendstoI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1880
  apply (drule (2) topological_tendstoD)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1881
  apply (simp only: eventually_sequentially)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1882
  apply (subgoal_tac "\<And>N k (n::nat). N + k <= n ==> N <= n - k")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1883
  apply metis
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1884
  by arith
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1885
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1886
lemma seq_offset_rev:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1887
  "((\<lambda>i. f(i + k)) ---> l) sequentially ==> (f ---> l) sequentially"
44584
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  1888
  by (rule LIMSEQ_offset) (* FIXME: redundant *)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1889
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1890
lemma seq_harmonic: "((\<lambda>n. inverse (real n)) ---> 0) sequentially"
44584
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  1891
  using LIMSEQ_inverse_real_of_nat by (rule LIMSEQ_imp_Suc)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1892
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1893
subsection {* More properties of closed balls *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1894
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1895
lemma closed_cball: "closed (cball x e)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1896
unfolding cball_def closed_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1897
unfolding Collect_neg_eq [symmetric] not_le
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1898
apply (clarsimp simp add: open_dist, rename_tac y)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1899
apply (rule_tac x="dist x y - e" in exI, clarsimp)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1900
apply (rename_tac x')
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1901
apply (cut_tac x=x and y=x' and z=y in dist_triangle)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1902
apply simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1903
done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1904
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1905
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
  1906
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1907
  { 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
  1908
    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
  1909
  } moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1910
  { 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
  1911
    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
  1912
  } ultimately
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1913
  show ?thesis unfolding open_contains_ball by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1914
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1915
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1916
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
  1917
  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
  1918
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1919
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
  1920
  apply (simp add: interior_def, safe)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1921
  apply (force simp add: open_contains_cball)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1922
  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
  1923
  apply (simp add: subset_trans [OF ball_subset_cball])
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1924
  done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1925
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1926
lemma islimpt_ball:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1927
  fixes x y :: "'a::{real_normed_vector,perfect_space}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1928
  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
  1929
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1930
  assume "?lhs"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1931
  { assume "e \<le> 0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1932
    hence *:"ball x e = {}" using ball_eq_empty[of x e] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1933
    have False using `?lhs` unfolding * using islimpt_EMPTY[of y] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1934
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1935
  hence "e > 0" by (metis not_less)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1936
  moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1937
  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
  1938
  ultimately show "?rhs" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1939
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1940
  assume "?rhs" hence "e>0"  by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1941
  { fix d::real assume "d>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1942
    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
  1943
    proof(cases "d \<le> dist x y")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1944
      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
  1945
      proof(cases "x=y")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1946
        case True hence False using `d \<le> dist x y` `d>0` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1947
        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
  1948
      next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1949
        case False
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1950
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1951
        have "dist x (y - (d / (2 * dist y x)) *\<^sub>R (y - x))
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1952
              = norm (x - y + (d / (2 * norm (y - x))) *\<^sub>R (y - x))"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1953
          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
  1954
        also have "\<dots> = \<bar>- 1 + d / (2 * norm (x - y))\<bar> * norm (x - y)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1955
          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
  1956
          unfolding scaleR_minus_left scaleR_one
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1957
          by (auto simp add: norm_minus_commute)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1958
        also have "\<dots> = \<bar>- norm (x - y) + d / 2\<bar>"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1959
          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
  1960
          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
  1961
        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
  1962
        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
  1963
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1964
        moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1965
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1966
        have "(d / (2*dist y x)) *\<^sub>R (y - x) \<noteq> 0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1967
          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
  1968
        moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1969
        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
  1970
          using `d>0` `x\<noteq>y`[unfolded dist_nz] dist_commute[of x y]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1971
          unfolding dist_norm by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1972
        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
  1973
      qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1974
    next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1975
      case False hence "d > dist x y" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1976
      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
  1977
      proof(cases "x=y")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1978
        case True
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1979
        obtain z where **: "z \<noteq> y" "dist z y < min e d"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1980
          using perfect_choose_dist[of "min e d" y]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1981
          using `d > 0` `e>0` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1982
        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
  1983
          unfolding `x = y`
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1984
          using `z \<noteq> y` **
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1985
          by (rule_tac x=z in bexI, auto simp add: dist_commute)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1986
      next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1987
        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
  1988
          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
  1989
      qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1990
    qed  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1991
  thus "?lhs" unfolding mem_cball islimpt_approachable mem_ball by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1992
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1993
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1994
lemma closure_ball_lemma:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1995
  fixes x y :: "'a::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1996
  assumes "x \<noteq> y" shows "y islimpt ball x (dist x y)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1997
proof (rule islimptI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1998
  fix T assume "y \<in> T" "open T"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1999
  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
  2000
    unfolding open_dist by fast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2001
  (* choose point between x and y, within distance r of y. *)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2002
  def k \<equiv> "min 1 (r / (2 * dist x y))"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2003
  def z \<equiv> "y + scaleR k (x - y)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2004
  have z_def2: "z = x + scaleR (1 - k) (y - x)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2005
    unfolding z_def by (simp add: algebra_simps)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2006
  have "dist z y < r"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2007
    unfolding z_def k_def using `0 < r`
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2008
    by (simp add: dist_norm min_def)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2009
  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
  2010
  have "dist x z < dist x y"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2011
    unfolding z_def2 dist_norm
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2012
    apply (simp add: norm_minus_commute)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2013
    apply (simp only: dist_norm [symmetric])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2014
    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
  2015
    apply (rule mult_strict_right_mono)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2016
    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
  2017
    apply (simp add: zero_less_dist_iff `x \<noteq> y`)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2018
    done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2019
  hence "z \<in> ball x (dist x y)" by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2020
  have "z \<noteq> y"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2021
    unfolding z_def k_def using `x \<noteq> y` `0 < r`
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2022
    by (simp add: min_def)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2023
  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
  2024
    using `z \<in> ball x (dist x y)` `z \<in> T` `z \<noteq> y`
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2025
    by fast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2026
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2027
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2028
lemma closure_ball:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2029
  fixes x :: "'a::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2030
  shows "0 < e \<Longrightarrow> closure (ball x e) = cball x e"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2031
apply (rule equalityI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2032
apply (rule closure_minimal)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2033
apply (rule ball_subset_cball)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2034
apply (rule closed_cball)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2035
apply (rule subsetI, rename_tac y)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2036
apply (simp add: le_less [where 'a=real])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2037
apply (erule disjE)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2038
apply (rule subsetD [OF closure_subset], simp)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2039
apply (simp add: closure_def)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2040
apply clarify
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2041
apply (rule closure_ball_lemma)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2042
apply (simp add: zero_less_dist_iff)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2043
done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2044
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2045
(* In a trivial vector space, this fails for e = 0. *)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2046
lemma interior_cball:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2047
  fixes x :: "'a::{real_normed_vector, perfect_space}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2048
  shows "interior (cball x e) = ball x e"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2049
proof(cases "e\<ge>0")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2050
  case False note cs = this
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2051
  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
  2052
  { fix y assume "y \<in> cball x e"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2053
    hence False unfolding mem_cball using dist_nz[of x y] cs by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2054
  hence "cball x e = {}" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2055
  hence "interior (cball x e) = {}" using interior_empty by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2056
  ultimately show ?thesis by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2057
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2058
  case True note cs = this
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2059
  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
  2060
  { 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
  2061
    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
  2062
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2063
    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
  2064
      using perfect_choose_dist [of d] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2065
    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
  2066
    hence xa_cball:"xa \<in> cball x e" using as(1) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2067
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2068
    hence "y \<in> ball x e" proof(cases "x = y")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2069
      case True
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2070
      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
  2071
      thus "y \<in> ball x e" using `x = y ` by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2072
    next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2073
      case False
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2074
      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
  2075
        using `d>0` norm_ge_zero[of "y - x"] `x \<noteq> y` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2076
      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
  2077
      have "y - x \<noteq> 0" using `x \<noteq> y` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2078
      hence **:"d / (2 * norm (y - x)) > 0" unfolding zero_less_norm_iff[THEN sym]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2079
        using `d>0` divide_pos_pos[of d "2*norm (y - x)"] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2080
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2081
      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
  2082
        by (auto simp add: dist_norm algebra_simps)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2083
      also have "\<dots> = norm ((1 + d / (2 * norm (y - x))) *\<^sub>R (y - x))"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2084
        by (auto simp add: algebra_simps)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2085
      also have "\<dots> = \<bar>1 + d / (2 * norm (y - x))\<bar> * norm (y - x)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2086
        using ** by auto
49962
a8cc904a6820 Renamed {left,right}_distrib to distrib_{right,left}.
webertj
parents: 49834
diff changeset
  2087
      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
  2088
      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
  2089
      thus "y \<in> ball x e" unfolding mem_ball using `d>0` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2090
    qed  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2091
  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
  2092
  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
  2093
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2094
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2095
lemma frontier_ball:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2096
  fixes a :: "'a::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2097
  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
  2098
  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
  2099
  apply (simp add: set_eq_iff)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2100
  by arith
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2101
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2102
lemma frontier_cball:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2103
  fixes a :: "'a::{real_normed_vector, perfect_space}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2104
  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
  2105
  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
  2106
  apply (simp add: set_eq_iff)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2107
  by arith
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2108
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2109
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
  2110
  apply (simp add: set_eq_iff not_le)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2111
  by (metis zero_le_dist dist_self order_less_le_trans)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2112
lemma cball_empty: "e < 0 ==> cball x e = {}" by (simp add: cball_eq_empty)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2113
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2114
lemma cball_eq_sing:
44072
5b970711fb39 class perfect_space inherits from topological_space;
huffman
parents: 43338
diff changeset
  2115
  fixes x :: "'a::{metric_space,perfect_space}"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2116
  shows "(cball x e = {x}) \<longleftrightarrow> e = 0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2117
proof (rule linorder_cases)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2118
  assume e: "0 < e"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2119
  obtain a where "a \<noteq> x" "dist a x < e"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2120
    using perfect_choose_dist [OF e] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2121
  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
  2122
  with e show ?thesis by (auto simp add: set_eq_iff)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2123
qed auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2124
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2125
lemma cball_sing:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2126
  fixes x :: "'a::metric_space"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2127
  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
  2128
  by (auto simp add: set_eq_iff)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2129
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  2130
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  2131
subsection {* Boundedness *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2132
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2133
  (* 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
  2134
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
  2135
  bounded :: "'a set \<Rightarrow> bool" where
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2136
  "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
  2137
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2138
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
  2139
unfolding bounded_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2140
apply safe
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2141
apply (rule_tac x="dist a x + e" in exI, clarify)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2142
apply (drule (1) bspec)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2143
apply (erule order_trans [OF dist_triangle add_left_mono])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2144
apply auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2145
done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2146
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2147
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
  2148
unfolding bounded_any_center [where a=0]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2149
by (simp add: dist_norm)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2150
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50094
diff changeset
  2151
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
  2152
  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
  2153
  using assms by auto
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50094
diff changeset
  2154
50948
8c742f9de9f5 generalize topology lemmas; simplify proofs
huffman
parents: 50944
diff changeset
  2155
lemma bounded_empty [simp]: "bounded {}"
8c742f9de9f5 generalize topology lemmas; simplify proofs
huffman
parents: 50944
diff changeset
  2156
  by (simp add: bounded_def)
8c742f9de9f5 generalize topology lemmas; simplify proofs
huffman
parents: 50944
diff changeset
  2157
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2158
lemma bounded_subset: "bounded T \<Longrightarrow> S \<subseteq> T ==> bounded S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2159
  by (metis bounded_def subset_eq)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2160
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2161
lemma bounded_interior[intro]: "bounded S ==> bounded(interior S)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2162
  by (metis bounded_subset interior_subset)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2163
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2164
lemma bounded_closure[intro]: assumes "bounded S" shows "bounded(closure S)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2165
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2166
  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
  2167
  { fix y assume "y \<in> closure S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2168
    then obtain f where f: "\<forall>n. f n \<in> S"  "(f ---> y) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2169
      unfolding closure_sequential by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2170
    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
  2171
    hence "eventually (\<lambda>n. dist x (f n) \<le> a) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2172
      by (rule eventually_mono, simp add: f(1))
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2173
    have "dist x y \<le> a"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2174
      apply (rule Lim_dist_ubound [of sequentially f])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2175
      apply (rule trivial_limit_sequentially)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2176
      apply (rule f(2))
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2177
      apply fact
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2178
      done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2179
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2180
  thus ?thesis unfolding bounded_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2181
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2182
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2183
lemma bounded_cball[simp,intro]: "bounded (cball x e)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2184
  apply (simp add: bounded_def)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2185
  apply (rule_tac x=x in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2186
  apply (rule_tac x=e in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2187
  apply auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2188
  done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2189
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2190
lemma bounded_ball[simp,intro]: "bounded(ball x e)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2191
  by (metis ball_subset_cball bounded_cball bounded_subset)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2192
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2193
lemma bounded_Un[simp]: "bounded (S \<union> T) \<longleftrightarrow> bounded S \<and> bounded T"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2194
  apply (auto simp add: bounded_def)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2195
  apply (rename_tac x y r s)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2196
  apply (rule_tac x=x in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2197
  apply (rule_tac x="max r (dist x y + s)" in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2198
  apply (rule ballI, rename_tac z, safe)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2199
  apply (drule (1) bspec, simp)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2200
  apply (drule (1) bspec)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2201
  apply (rule min_max.le_supI2)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2202
  apply (erule order_trans [OF dist_triangle add_left_mono])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2203
  done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2204
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2205
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
  2206
  by (induct rule: finite_induct[of F], auto)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2207
50955
ada575c605e1 simplify proof of compact_imp_bounded
huffman
parents: 50949
diff changeset
  2208
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
  2209
  by (induct set: finite, auto)
ada575c605e1 simplify proof of compact_imp_bounded
huffman
parents: 50949
diff changeset
  2210
50948
8c742f9de9f5 generalize topology lemmas; simplify proofs
huffman
parents: 50944
diff changeset
  2211
lemma bounded_insert [simp]: "bounded (insert x S) \<longleftrightarrow> bounded S"
8c742f9de9f5 generalize topology lemmas; simplify proofs
huffman
parents: 50944
diff changeset
  2212
proof -
8c742f9de9f5 generalize topology lemmas; simplify proofs
huffman
parents: 50944
diff changeset
  2213
  have "\<forall>y\<in>{x}. dist x y \<le> 0" by simp
8c742f9de9f5 generalize topology lemmas; simplify proofs
huffman
parents: 50944
diff changeset
  2214
  hence "bounded {x}" unfolding bounded_def by fast
8c742f9de9f5 generalize topology lemmas; simplify proofs
huffman
parents: 50944
diff changeset
  2215
  thus ?thesis by (metis insert_is_Un bounded_Un)
8c742f9de9f5 generalize topology lemmas; simplify proofs
huffman
parents: 50944
diff changeset
  2216
qed
8c742f9de9f5 generalize topology lemmas; simplify proofs
huffman
parents: 50944
diff changeset
  2217
8c742f9de9f5 generalize topology lemmas; simplify proofs
huffman
parents: 50944
diff changeset
  2218
lemma finite_imp_bounded [intro]: "finite S \<Longrightarrow> bounded S"
8c742f9de9f5 generalize topology lemmas; simplify proofs
huffman
parents: 50944
diff changeset
  2219
  by (induct set: finite, simp_all)
8c742f9de9f5 generalize topology lemmas; simplify proofs
huffman
parents: 50944
diff changeset
  2220
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2221
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
  2222
  apply (simp add: bounded_iff)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2223
  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
  2224
  by metis arith
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2225
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2226
lemma bounded_Int[intro]: "bounded S \<or> bounded T \<Longrightarrow> bounded (S \<inter> T)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2227
  by (metis Int_lower1 Int_lower2 bounded_subset)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2228
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2229
lemma bounded_diff[intro]: "bounded S ==> bounded (S - T)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2230
apply (metis Diff_subset bounded_subset)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2231
done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2232
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2233
lemma not_bounded_UNIV[simp, intro]:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2234
  "\<not> bounded (UNIV :: 'a::{real_normed_vector, perfect_space} set)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2235
proof(auto simp add: bounded_pos not_le)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2236
  obtain x :: 'a where "x \<noteq> 0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2237
    using perfect_choose_dist [OF zero_less_one] by fast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2238
  fix b::real  assume b: "b >0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2239
  have b1: "b +1 \<ge> 0" using b by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2240
  with `x \<noteq> 0` have "b < norm (scaleR (b + 1) (sgn x))"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2241
    by (simp add: norm_sgn)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2242
  then show "\<exists>x::'a. b < norm x" ..
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2243
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2244
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2245
lemma bounded_linear_image:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2246
  assumes "bounded S" "bounded_linear f"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2247
  shows "bounded(f ` S)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2248
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2249
  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
  2250
  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
  2251
  { fix x assume "x\<in>S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2252
    hence "norm x \<le> b" using b by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2253
    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
  2254
      by (metis B(1) B(2) order_trans mult_le_cancel_left_pos)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2255
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2256
  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
  2257
    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
  2258
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2259
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2260
lemma bounded_scaling:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2261
  fixes S :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2262
  shows "bounded S \<Longrightarrow> bounded ((\<lambda>x. c *\<^sub>R x) ` S)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2263
  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
  2264
  apply (rule bounded_linear_scaleR_right)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2265
  done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2266
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2267
lemma bounded_translation:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2268
  fixes S :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2269
  assumes "bounded S" shows "bounded ((\<lambda>x. a + x) ` S)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2270
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2271
  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
  2272
  { fix x assume "x\<in>S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2273
    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
  2274
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2275
  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
  2276
    by (auto intro!: exI[of _ "b + norm a"])
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2277
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2278
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2279
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2280
text{* Some theorems on sups and infs using the notion "bounded". *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2281
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2282
lemma bounded_real:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2283
  fixes S :: "real set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2284
  shows "bounded S \<longleftrightarrow>  (\<exists>a. \<forall>x\<in>S. abs x <= a)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2285
  by (simp add: bounded_iff)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2286
33270
paulson
parents: 33175
diff changeset
  2287
lemma bounded_has_Sup:
paulson
parents: 33175
diff changeset
  2288
  fixes S :: "real set"
paulson
parents: 33175
diff changeset
  2289
  assumes "bounded S" "S \<noteq> {}"
paulson
parents: 33175
diff changeset
  2290
  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
  2291
proof
paulson
parents: 33175
diff changeset
  2292
  fix x assume "x\<in>S"
paulson
parents: 33175
diff changeset
  2293
  thus "x \<le> Sup S"
paulson
parents: 33175
diff changeset
  2294
    by (metis SupInf.Sup_upper abs_le_D1 assms(1) bounded_real)
paulson
parents: 33175
diff changeset
  2295
next
paulson
parents: 33175
diff changeset
  2296
  show "\<forall>b. (\<forall>x\<in>S. x \<le> b) \<longrightarrow> Sup S \<le> b" using assms
paulson
parents: 33175
diff changeset
  2297
    by (metis SupInf.Sup_least)
paulson
parents: 33175
diff changeset
  2298
qed
paulson
parents: 33175
diff changeset
  2299
paulson
parents: 33175
diff changeset
  2300
lemma Sup_insert:
paulson
parents: 33175
diff changeset
  2301
  fixes S :: "real set"
paulson
parents: 33175
diff changeset
  2302
  shows "bounded S ==> Sup(insert x S) = (if S = {} then x else max x (Sup S))" 
paulson
parents: 33175
diff changeset
  2303
by auto (metis Int_absorb Sup_insert_nonempty assms bounded_has_Sup(1) disjoint_iff_not_equal) 
paulson
parents: 33175
diff changeset
  2304
paulson
parents: 33175
diff changeset
  2305
lemma Sup_insert_finite:
paulson
parents: 33175
diff changeset
  2306
  fixes S :: "real set"
paulson
parents: 33175
diff changeset
  2307
  shows "finite S \<Longrightarrow> Sup(insert x S) = (if S = {} then x else max x (Sup S))"
paulson
parents: 33175
diff changeset
  2308
  apply (rule Sup_insert)
paulson
parents: 33175
diff changeset
  2309
  apply (rule finite_imp_bounded)
paulson
parents: 33175
diff changeset
  2310
  by simp
paulson
parents: 33175
diff changeset
  2311
paulson
parents: 33175
diff changeset
  2312
lemma bounded_has_Inf:
paulson
parents: 33175
diff changeset
  2313
  fixes S :: "real set"
paulson
parents: 33175
diff changeset
  2314
  assumes "bounded S"  "S \<noteq> {}"
paulson
parents: 33175
diff changeset
  2315
  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
  2316
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2317
  fix x assume "x\<in>S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2318
  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
  2319
  thus "x \<ge> Inf S" using `x\<in>S`
paulson
parents: 33175
diff changeset
  2320
    by (metis Inf_lower_EX abs_le_D2 minus_le_iff)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2321
next
33270
paulson
parents: 33175
diff changeset
  2322
  show "\<forall>b. (\<forall>x\<in>S. x >= b) \<longrightarrow> Inf S \<ge> b" using assms
paulson
parents: 33175
diff changeset
  2323
    by (metis SupInf.Inf_greatest)
paulson
parents: 33175
diff changeset
  2324
qed
paulson
parents: 33175
diff changeset
  2325
paulson
parents: 33175
diff changeset
  2326
lemma Inf_insert:
paulson
parents: 33175
diff changeset
  2327
  fixes S :: "real set"
paulson
parents: 33175
diff changeset
  2328
  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
  2329
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
  2330
33270
paulson
parents: 33175
diff changeset
  2331
lemma Inf_insert_finite:
paulson
parents: 33175
diff changeset
  2332
  fixes S :: "real set"
paulson
parents: 33175
diff changeset
  2333
  shows "finite S ==> Inf(insert x S) = (if S = {} then x else min x (Inf S))"
paulson
parents: 33175
diff changeset
  2334
  by (rule Inf_insert, rule finite_imp_bounded, simp)
paulson
parents: 33175
diff changeset
  2335
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2336
subsection {* Compactness *}
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2337
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2338
subsubsection{* Open-cover compactness *}
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2339
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2340
definition compact :: "'a::topological_space set \<Rightarrow> bool" where
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2341
  compact_eq_heine_borel: -- "This name is used for backwards compatibility"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2342
    "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
  2343
50898
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2344
lemma compactI:
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2345
  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
  2346
  shows "compact s"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2347
  unfolding compact_eq_heine_borel using assms by metis
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2348
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2349
lemma compactE:
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2350
  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
  2351
  obtains C' where "C' \<subseteq> C" and "finite C'" and "s \<subseteq> \<Union>C'"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2352
  using assms unfolding compact_eq_heine_borel by metis
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2353
50944
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  2354
lemma compactE_image:
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  2355
  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
  2356
  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
  2357
  using assms unfolding ball_simps[symmetric] SUP_def
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  2358
  by (metis (lifting) finite_subset_image compact_eq_heine_borel[of s])
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  2359
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2360
subsubsection {* Bolzano-Weierstrass property *}
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2361
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2362
lemma heine_borel_imp_bolzano_weierstrass:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2363
  assumes "compact s" "infinite t"  "t \<subseteq> s"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2364
  shows "\<exists>x \<in> s. x islimpt t"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2365
proof(rule ccontr)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2366
  assume "\<not> (\<exists>x \<in> s. x islimpt t)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2367
  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
  2368
    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
  2369
  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
  2370
    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
  2371
  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
  2372
  { 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
  2373
    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
  2374
    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
  2375
  hence "inj_on f t" unfolding inj_on_def by simp
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2376
  hence "infinite (f ` t)" using assms(2) using finite_imageD by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2377
  moreover
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2378
  { fix x assume "x\<in>t" "f x \<notin> g"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2379
    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
  2380
    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
  2381
    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
  2382
    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
  2383
  hence "f ` t \<subseteq> g" by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2384
  ultimately show False using g(2) using finite_subset by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2385
qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2386
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
  2387
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
  2388
  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
  2389
  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
  2390
  shows "\<exists>r. subseq r \<and> (f \<circ> r) ----> l"
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2391
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
  2392
  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
  2393
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2394
  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
  2395
  { 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
  2396
    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
  2397
      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
  2398
    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
  2399
      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
  2400
    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
  2401
      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
  2402
    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
  2403
      by (auto simp: not_le)
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2404
    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
  2405
      unfolding s_def by (auto intro: someI2_ex) }
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2406
  note s = this
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2407
  def r \<equiv> "nat_rec (s 0 0) s"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2408
  have "subseq r"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2409
    by (auto simp: r_def s subseq_Suc_iff)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2410
  moreover
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2411
  have "(\<lambda>n. f (r n)) ----> l"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2412
  proof (rule topological_tendstoI)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2413
    fix S assume "open S" "l \<in> S"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2414
    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
  2415
    moreover
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2416
    { 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
  2417
        by (cases i) (simp_all add: r_def s) }
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2418
    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
  2419
    ultimately show "eventually (\<lambda>i. f (r i) \<in> S) sequentially"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2420
      by eventually_elim auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2421
  qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2422
  ultimately show "\<exists>r. subseq r \<and> (f \<circ> r) ----> l"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2423
    by (auto simp: convergent_def comp_def)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2424
qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2425
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2426
lemma sequence_infinite_lemma:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2427
  fixes f :: "nat \<Rightarrow> 'a::t1_space"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2428
  assumes "\<forall>n. f n \<noteq> l" and "(f ---> l) sequentially"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2429
  shows "infinite (range f)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2430
proof
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2431
  assume "finite (range f)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2432
  hence "closed (range f)" by (rule finite_imp_closed)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2433
  hence "open (- range f)" by (rule open_Compl)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2434
  from assms(1) have "l \<in> - range f" by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2435
  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
  2436
    using `open (- range f)` `l \<in> - range f` by (rule topological_tendstoD)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2437
  thus False unfolding eventually_sequentially by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2438
qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2439
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2440
lemma closure_insert:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2441
  fixes x :: "'a::t1_space"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2442
  shows "closure (insert x s) = insert x (closure s)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2443
apply (rule closure_unique)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2444
apply (rule insert_mono [OF closure_subset])
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2445
apply (rule closed_insert [OF closed_closure])
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2446
apply (simp add: closure_minimal)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2447
done
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2448
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2449
lemma islimpt_insert:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2450
  fixes x :: "'a::t1_space"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2451
  shows "x islimpt (insert a s) \<longleftrightarrow> x islimpt s"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2452
proof
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2453
  assume *: "x islimpt (insert a s)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2454
  show "x islimpt s"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2455
  proof (rule islimptI)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2456
    fix t assume t: "x \<in> t" "open t"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2457
    show "\<exists>y\<in>s. y \<in> t \<and> y \<noteq> x"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2458
    proof (cases "x = a")
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2459
      case True
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2460
      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
  2461
        using * t by (rule islimptE)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2462
      with `x = a` show ?thesis by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2463
    next
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2464
      case False
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2465
      with t have t': "x \<in> t - {a}" "open (t - {a})"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2466
        by (simp_all add: open_Diff)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2467
      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
  2468
        using * t' by (rule islimptE)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2469
      thus ?thesis by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2470
    qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2471
  qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2472
next
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2473
  assume "x islimpt s" thus "x islimpt (insert a s)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2474
    by (rule islimpt_subset) auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2475
qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2476
50897
078590669527 generalize lemma islimpt_finite to class t1_space
huffman
parents: 50884
diff changeset
  2477
lemma islimpt_finite:
078590669527 generalize lemma islimpt_finite to class t1_space
huffman
parents: 50884
diff changeset
  2478
  fixes x :: "'a::t1_space"
078590669527 generalize lemma islimpt_finite to class t1_space
huffman
parents: 50884
diff changeset
  2479
  shows "finite s \<Longrightarrow> \<not> x islimpt s"
078590669527 generalize lemma islimpt_finite to class t1_space
huffman
parents: 50884
diff changeset
  2480
by (induct set: finite, simp_all add: islimpt_insert)
078590669527 generalize lemma islimpt_finite to class t1_space
huffman
parents: 50884
diff changeset
  2481
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2482
lemma islimpt_union_finite:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2483
  fixes x :: "'a::t1_space"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2484
  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
  2485
by (simp add: islimpt_Un islimpt_finite)
078590669527 generalize lemma islimpt_finite to class t1_space
huffman
parents: 50884
diff changeset
  2486
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
  2487
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
  2488
  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
  2489
  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
  2490
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
  2491
  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
  2492
  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
  2493
    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
  2494
  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
  2495
    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
  2496
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
  2497
  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
  2498
  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
  2499
  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
  2500
    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
  2501
  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
  2502
    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
  2503
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
  2504
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
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
  2506
  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
  2507
  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
  2508
  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
  2509
  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
  2510
  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
  2511
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2512
lemma sequence_unique_limpt:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2513
  fixes f :: "nat \<Rightarrow> 'a::t2_space"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2514
  assumes "(f ---> l) sequentially" and "l' islimpt (range f)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2515
  shows "l' = l"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2516
proof (rule ccontr)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2517
  assume "l' \<noteq> l"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2518
  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
  2519
    using hausdorff [OF `l' \<noteq> l`] by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2520
  have "eventually (\<lambda>n. f n \<in> t) sequentially"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2521
    using assms(1) `open t` `l \<in> t` by (rule topological_tendstoD)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2522
  then obtain N where "\<forall>n\<ge>N. f n \<in> t"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2523
    unfolding eventually_sequentially by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2524
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2525
  have "UNIV = {..<N} \<union> {N..}" by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2526
  hence "l' islimpt (f ` ({..<N} \<union> {N..}))" using assms(2) by simp
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2527
  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
  2528
  hence "l' islimpt (f ` {N..})" by (simp add: islimpt_union_finite)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2529
  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
  2530
    using `l' \<in> s` `open s` by (rule islimptE)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2531
  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
  2532
  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
  2533
  with `s \<inter> t = {}` show False by simp
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2534
qed
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
lemma bolzano_weierstrass_imp_closed:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2537
  fixes s :: "'a::{first_countable_topology, t2_space} set"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2538
  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
  2539
  shows "closed s"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2540
proof-
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2541
  { 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
  2542
    hence "l \<in> s"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2543
    proof(cases "\<forall>n. x n \<noteq> l")
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2544
      case False thus "l\<in>s" using as(1) by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2545
    next
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2546
      case True note cas = this
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2547
      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
  2548
      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
  2549
      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
  2550
    qed  }
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2551
  thus ?thesis unfolding closed_sequential_limits by fast
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2552
qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2553
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2554
lemma compact_imp_closed:
50898
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2555
  fixes s :: "'a::t2_space set"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2556
  assumes "compact s" shows "closed s"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2557
unfolding closed_def
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2558
proof (rule openI)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2559
  fix y assume "y \<in> - s"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2560
  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
  2561
  note `compact s`
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2562
  moreover have "\<forall>u\<in>?C. open u" by simp
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2563
  moreover have "s \<subseteq> \<Union>?C"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2564
  proof
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2565
    fix x assume "x \<in> s"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2566
    with `y \<in> - s` have "x \<noteq> y" by clarsimp
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2567
    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
  2568
      by (rule hausdorff)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2569
    with `x \<in> s` show "x \<in> \<Union>?C"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2570
      unfolding eventually_nhds by auto
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2571
  qed
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2572
  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
  2573
    by (rule compactE)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2574
  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
  2575
  with `finite D` have "eventually (\<lambda>y. y \<notin> \<Union>D) (nhds y)"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2576
    by (simp add: eventually_Ball_finite)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2577
  with `s \<subseteq> \<Union>D` have "eventually (\<lambda>y. y \<notin> s) (nhds y)"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2578
    by (auto elim!: eventually_mono [rotated])
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2579
  thus "\<exists>t. open t \<and> y \<in> t \<and> t \<subseteq> - s"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2580
    by (simp add: eventually_nhds subset_eq)
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2581
qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2582
50944
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  2583
lemma compact_imp_bounded:
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  2584
  assumes "compact U" shows "bounded U"
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  2585
proof -
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  2586
  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
  2587
  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
  2588
    by (elim compactE_image)
50955
ada575c605e1 simplify proof of compact_imp_bounded
huffman
parents: 50949
diff changeset
  2589
  from `finite D` have "bounded (\<Union>x\<in>D. ball x 1)"
ada575c605e1 simplify proof of compact_imp_bounded
huffman
parents: 50949
diff changeset
  2590
    by (simp add: bounded_UN)
ada575c605e1 simplify proof of compact_imp_bounded
huffman
parents: 50949
diff changeset
  2591
  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
  2592
    by (rule bounded_subset)
50944
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  2593
qed
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  2594
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2595
text{* In particular, some common special cases. *}
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2596
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2597
lemma compact_empty[simp]:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2598
 "compact {}"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2599
  unfolding compact_eq_heine_borel
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2600
  by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2601
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2602
lemma compact_union [intro]:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2603
  assumes "compact s" "compact t" shows " compact (s \<union> t)"
50898
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2604
proof (rule compactI)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2605
  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
  2606
  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
  2607
    unfolding compact_eq_heine_borel by (auto elim!: allE[of _ f]) metis
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2608
  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
  2609
    unfolding compact_eq_heine_borel by (auto elim!: allE[of _ f]) metis
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2610
  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
  2611
    by (auto intro!: exI[of _ "s' \<union> t'"])
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2612
qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2613
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2614
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
  2615
  by (induct set: finite) auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2616
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2617
lemma compact_UN [intro]:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2618
  "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
  2619
  unfolding SUP_def by (rule compact_Union) auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2620
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2621
lemma compact_inter_closed [intro]:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2622
  assumes "compact s" and "closed t"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2623
  shows "compact (s \<inter> t)"
50898
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2624
proof (rule compactI)
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2625
  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
  2626
  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
  2627
  moreover from cover have "s \<subseteq> \<Union>(C \<union> {-t})" by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2628
  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
  2629
    using `compact s` unfolding compact_eq_heine_borel by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2630
  then guess D ..
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2631
  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
  2632
    by (intro exI[of _ "D - {-t}"]) auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2633
qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2634
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2635
lemma closed_inter_compact [intro]:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2636
  assumes "closed s" and "compact t"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2637
  shows "compact (s \<inter> t)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2638
  using compact_inter_closed [of t s] assms
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2639
  by (simp add: Int_commute)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2640
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2641
lemma compact_inter [intro]:
50898
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2642
  fixes s t :: "'a :: t2_space set"
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2643
  assumes "compact s" and "compact t"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2644
  shows "compact (s \<inter> t)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2645
  using assms by (intro compact_inter_closed compact_imp_closed)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2646
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2647
lemma compact_sing [simp]: "compact {a}"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2648
  unfolding compact_eq_heine_borel by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2649
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2650
lemma compact_insert [simp]:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2651
  assumes "compact s" shows "compact (insert x s)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2652
proof -
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2653
  have "compact ({x} \<union> s)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2654
    using compact_sing assms by (rule compact_union)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2655
  thus ?thesis by simp
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2656
qed
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 finite_imp_compact:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2659
  shows "finite s \<Longrightarrow> compact s"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2660
  by (induct set: finite) simp_all
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2661
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2662
lemma open_delete:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2663
  fixes s :: "'a::t1_space set"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2664
  shows "open s \<Longrightarrow> open (s - {x})"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2665
  by (simp add: open_Diff)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2666
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2667
text{* Finite intersection property *}
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 inj_setminus: "inj_on uminus (A::'a set set)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2670
  by (auto simp: inj_on_def)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2671
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2672
lemma compact_fip:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2673
  "compact U \<longleftrightarrow>
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2674
    (\<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
  2675
  (is "_ \<longleftrightarrow> ?R")
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2676
proof (safe intro!: compact_eq_heine_borel[THEN iffD2])
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2677
  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
  2678
    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
  2679
  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
  2680
    by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2681
  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
  2682
    unfolding compact_eq_heine_borel by (metis subset_image_iff)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2683
  with fi[THEN spec, of B] show False
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2684
    by (auto dest: finite_imageD intro: inj_setminus)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2685
next
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2686
  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
  2687
  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
  2688
    by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2689
  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
  2690
    by (metis subset_image_iff)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2691
  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
  2692
    by  (auto intro!: exI[of _ B] inj_setminus dest: finite_imageD)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2693
qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2694
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2695
lemma compact_imp_fip:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2696
  "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
  2697
    s \<inter> (\<Inter> f) \<noteq> {}"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2698
  unfolding compact_fip by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2699
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2700
text{*Compactness expressed with filters*}
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2701
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2702
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
  2703
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2704
lemma eventually_filter_from_subbase:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2705
  "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
  2706
    (is "_ \<longleftrightarrow> ?R P")
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2707
  unfolding filter_from_subbase_def
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2708
proof (rule eventually_Abs_filter is_filter.intro)+
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2709
  show "?R (\<lambda>x. True)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2710
    by (rule exI[of _ "{}"]) (simp add: le_fun_def)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2711
next
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2712
  fix P Q assume "?R P" then guess X ..
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2713
  moreover assume "?R Q" then guess Y ..
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2714
  ultimately show "?R (\<lambda>x. P x \<and> Q x)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2715
    by (intro exI[of _ "X \<union> Y"]) auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2716
next
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2717
  fix P Q
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2718
  assume "?R P" then guess X ..
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2719
  moreover assume "\<forall>x. P x \<longrightarrow> Q x"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2720
  ultimately show "?R Q"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2721
    by (intro exI[of _ X]) auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2722
qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2723
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2724
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
  2725
  by (subst eventually_filter_from_subbase) (auto intro!: exI[of _ "{P}"])
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2726
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2727
lemma filter_from_subbase_not_bot:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2728
  "\<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
  2729
  unfolding trivial_limit_def eventually_filter_from_subbase by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2730
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2731
lemma closure_iff_nhds_not_empty:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2732
  "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
  2733
proof safe
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2734
  assume x: "x \<in> closure X"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2735
  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
  2736
  then have "x \<notin> closure (-S)" 
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2737
    by (auto simp: closure_complement subset_eq[symmetric] intro: interiorI)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2738
  with x have "x \<in> closure X - closure (-S)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2739
    by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2740
  also have "\<dots> \<subseteq> closure (X \<inter> S)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2741
    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
  2742
  finally have "X \<inter> S \<noteq> {}" by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2743
  then show False using `X \<inter> A = {}` `S \<subseteq> A` by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2744
next
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2745
  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
  2746
  from this[THEN spec, of "- X", THEN spec, of "- closure X"]
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2747
  show "x \<in> closure X"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2748
    by (simp add: closure_subset open_Compl)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2749
qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2750
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2751
lemma compact_filter:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2752
  "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
  2753
proof (intro allI iffI impI compact_fip[THEN iffD2] notI)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2754
  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
  2755
  from F have "U \<noteq> {}"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2756
    by (auto simp: eventually_False)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2757
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2758
  def Z \<equiv> "closure ` {A. eventually (\<lambda>x. x \<in> A) F}"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2759
  then have "\<forall>z\<in>Z. closed z"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2760
    by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2761
  moreover 
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2762
  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
  2763
    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
  2764
  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
  2765
  proof (intro allI impI)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2766
    fix B assume "finite B" "B \<subseteq> Z"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2767
    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
  2768
      by (auto intro!: eventually_Ball_finite)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2769
    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
  2770
      by eventually_elim auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2771
    with F show "U \<inter> \<Inter>B \<noteq> {}"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2772
      by (intro notI) (simp add: eventually_False)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2773
  qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2774
  ultimately have "U \<inter> \<Inter>Z \<noteq> {}"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2775
    using `compact U` unfolding compact_fip by blast
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2776
  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
  2777
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2778
  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
  2779
    unfolding eventually_inf eventually_nhds
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2780
  proof safe
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2781
    fix P Q R S
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2782
    assume "eventually R F" "open S" "x \<in> S"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2783
    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
  2784
    have "S \<inter> {x. R x} \<noteq> {}" by (auto simp: Z_def)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2785
    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
  2786
    ultimately show False by (auto simp: set_eq_iff)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2787
  qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2788
  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
  2789
    by (metis eventually_bot)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2790
next
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2791
  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
  2792
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2793
  def P' \<equiv> "(\<lambda>a (x::'a). x \<in> a)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2794
  then have inj_P': "\<And>A. inj_on P' A"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2795
    by (auto intro!: inj_onI simp: fun_eq_iff)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2796
  def F \<equiv> "filter_from_subbase (P' ` insert U A)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2797
  have "F \<noteq> bot"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2798
    unfolding F_def
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2799
  proof (safe intro!: filter_from_subbase_not_bot)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2800
    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
  2801
    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
  2802
      unfolding subset_image_iff by (auto intro: inj_P' finite_imageD)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2803
    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
  2804
    with B show False by (auto simp: P'_def fun_eq_iff)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2805
  qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2806
  moreover have "eventually (\<lambda>x. x \<in> U) F"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2807
    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
  2808
  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
  2809
  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
  2810
    by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2811
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2812
  { fix V assume "V \<in> A"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2813
    then have V: "eventually (\<lambda>x. x \<in> V) F"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2814
      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
  2815
    have "x \<in> closure V"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2816
      unfolding closure_iff_nhds_not_empty
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2817
    proof (intro impI allI)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2818
      fix S A assume "open S" "x \<in> S" "S \<subseteq> A"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2819
      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
  2820
      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
  2821
        by (auto simp: eventually_inf)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2822
      with x show "V \<inter> A \<noteq> {}"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2823
        by (auto simp del: Int_iff simp add: trivial_limit_def)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2824
    qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2825
    then have "x \<in> V"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2826
      using `V \<in> A` A(1) by simp }
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2827
  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
  2828
  with `U \<inter> \<Inter>A = {}` show False by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2829
qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2830
50941
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2831
definition "countably_compact U \<longleftrightarrow>
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2832
    (\<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
  2833
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2834
lemma countably_compactE:
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2835
  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
  2836
  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
  2837
  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
  2838
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2839
lemma countably_compactI:
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2840
  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
  2841
  shows "countably_compact s"
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2842
  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
  2843
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2844
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
  2845
  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
  2846
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2847
lemma countably_compact_imp_compact:
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2848
  assumes "countably_compact U"
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2849
  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
  2850
  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
  2851
  shows "compact U"
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2852
  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
  2853
proof safe
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2854
  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
  2855
  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
  2856
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2857
  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
  2858
  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
  2859
    unfolding C_def using ccover by auto
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2860
  moreover
50941
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2861
  have "\<Union>A \<inter> U \<subseteq> \<Union>C"
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2862
  proof safe
50941
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2863
    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
  2864
    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
  2865
    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
  2866
      by auto
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2867
  qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2868
  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
  2869
  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
  2870
    using * by metis
50941
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2871
  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
  2872
    by (auto simp: C_def)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2873
  then guess f unfolding bchoice_iff Bex_def ..
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2874
  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
  2875
    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
  2876
qed
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2877
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2878
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
  2879
  "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
  2880
proof (rule countably_compact_imp_compact)
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2881
  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
  2882
  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
  2883
  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
  2884
qed (insert countable_basis topological_basis_open[OF is_basis], auto)
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  2885
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
  2886
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
  2887
  "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
  2888
  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
  2889
  
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  2890
subsubsection{* Sequential compactness *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2891
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2892
definition
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2893
  seq_compact :: "'a::topological_space set \<Rightarrow> bool" where
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2894
  "seq_compact S \<longleftrightarrow>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2895
   (\<forall>f. (\<forall>n. f n \<in> S) \<longrightarrow>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2896
       (\<exists>l\<in>S. \<exists>r. subseq r \<and> ((f o r) ---> l) sequentially))"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2897
50941
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2898
lemma seq_compact_imp_countably_compact:
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2899
  fixes U :: "'a :: first_countable_topology set"
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2900
  assumes "seq_compact U"
50941
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2901
  shows "countably_compact U"
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  2902
proof (safe intro!: countably_compactI)
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2903
  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
  2904
  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
  2905
    using `seq_compact U` by (fastforce simp: seq_compact_def subset_eq)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2906
  show "\<exists>T\<subseteq>A. finite T \<and> U \<subseteq> \<Union>T"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2907
  proof cases
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2908
    assume "finite A" with A show ?thesis by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2909
  next
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2910
    assume "infinite A"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2911
    then have "A \<noteq> {}" by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2912
    show ?thesis
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2913
    proof (rule ccontr)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2914
      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
  2915
      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
  2916
      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
  2917
      def X \<equiv> "\<lambda>n. X' (from_nat_into A ` {.. n})"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2918
      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
  2919
        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
  2920
      then have "range X \<subseteq> U" by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2921
      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
  2922
      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
  2923
      obtain n where "x \<in> from_nat_into A n" by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2924
      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
  2925
      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
  2926
        unfolding tendsto_def by (auto simp: comp_def)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2927
      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
  2928
        by (auto simp: eventually_sequentially)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2929
      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
  2930
        by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2931
      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
  2932
        by (auto intro!: exI[of _ "max n N"])
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2933
      ultimately show False
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2934
        by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2935
    qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2936
  qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2937
qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2938
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2939
lemma compact_imp_seq_compact:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2940
  fixes U :: "'a :: first_countable_topology set"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2941
  assumes "compact U" shows "seq_compact U"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2942
  unfolding seq_compact_def
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2943
proof safe
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2944
  fix X :: "nat \<Rightarrow> 'a" assume "\<forall>n. X n \<in> U"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2945
  then have "eventually (\<lambda>x. x \<in> U) (filtermap X sequentially)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2946
    by (auto simp: eventually_filtermap)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2947
  moreover have "filtermap X sequentially \<noteq> bot"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2948
    by (simp add: trivial_limit_def eventually_filtermap)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2949
  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
  2950
    using `compact U` by (auto simp: compact_filter)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2951
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2952
  from countable_basis_at_decseq[of x] guess A . note A = this
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2953
  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
  2954
  { fix n i
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2955
    have "\<exists>a. i < a \<and> X a \<in> A (Suc n)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2956
    proof (rule ccontr)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2957
      assume "\<not> (\<exists>a>i. X a \<in> A (Suc n))"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2958
      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
  2959
      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
  2960
        by (auto simp: eventually_filtermap eventually_sequentially)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2961
      moreover have "eventually (\<lambda>x. x \<in> A (Suc n)) (nhds x)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2962
        using A(1,2)[of "Suc n"] by (auto simp: eventually_nhds)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2963
      ultimately have "eventually (\<lambda>x. False) ?F"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2964
        by (auto simp add: eventually_inf)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2965
      with x show False
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2966
        by (simp add: eventually_False)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2967
    qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2968
    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
  2969
      unfolding s_def by (auto intro: someI2_ex) }
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2970
  note s = this
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2971
  def r \<equiv> "nat_rec (s 0 0) s"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2972
  have "subseq r"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2973
    by (auto simp: r_def s subseq_Suc_iff)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2974
  moreover
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2975
  have "(\<lambda>n. X (r n)) ----> x"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2976
  proof (rule topological_tendstoI)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2977
    fix S assume "open S" "x \<in> S"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2978
    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
  2979
    moreover
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2980
    { 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
  2981
        by (cases i) (simp_all add: r_def s) }
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2982
    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
  2983
    ultimately show "eventually (\<lambda>i. X (r i) \<in> S) sequentially"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2984
      by eventually_elim auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2985
  qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2986
  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
  2987
    using `x \<in> U` by (auto simp: convergent_def comp_def)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2988
qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2989
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2990
lemma seq_compactI:
44075
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  2991
  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
  2992
  shows "seq_compact S"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2993
  unfolding seq_compact_def using assms by fast
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2994
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2995
lemma seq_compactE:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2996
  assumes "seq_compact S" "\<forall>n. f n \<in> S"
44075
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  2997
  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
  2998
  using assms unfolding seq_compact_def by fast
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2999
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
  3000
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
  3001
  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
  3002
  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
  3003
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
  3004
  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
  3005
  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
  3006
  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
  3007
    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
  3008
  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
  3009
  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
  3010
  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
  3011
  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
  3012
    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
  3013
    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
  3014
    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
  3015
    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
  3016
    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
  3017
    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
  3018
  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
  3019
  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
  3020
    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
  3021
  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
  3022
  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
  3023
    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
  3024
    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
  3025
  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
  3026
    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
  3027
  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
  3028
    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
  3029
  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
  3030
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
  3031
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
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
  3033
  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
  3034
  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
  3035
  shows "seq_compact s"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  3036
proof -
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  3037
  { 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
  3038
    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
  3039
    proof (cases "finite (range f)")
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  3040
      case True
50941
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  3041
      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
  3042
        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
  3043
      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
  3044
        using infinite_enumerate by blast
3690724028b1 add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents: 50940
diff changeset
  3045
      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
  3046
        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
  3047
      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
  3048
        by auto
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  3049
    next
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  3050
      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
  3051
      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
  3052
      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
  3053
      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
  3054
        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
  3055
      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
  3056
    qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  3057
  }
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  3058
  thus ?thesis unfolding seq_compact_def by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  3059
qed
44075
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3060
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
  3061
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
  3062
  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
  3063
  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
  3064
  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
  3065
    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
  3066
    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
  3067
    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
  3068
  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
  3069
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
  3070
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
  3071
  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
  3072
  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
  3073
  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
  3074
    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
  3075
    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
  3076
    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
  3077
  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
  3078
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
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
  3080
  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
  3081
  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
  3082
  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
  3083
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
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
  3085
  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
  3086
  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
  3087
  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
  3088
50940
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3089
subsubsection{* Total boundedness *}
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3090
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3091
lemma cauchy_def:
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3092
  "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
  3093
unfolding Cauchy_def by blast
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3094
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3095
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
  3096
  "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
  3097
declare helper_1.simps[simp del]
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3098
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3099
lemma seq_compact_imp_totally_bounded:
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3100
  assumes "seq_compact s"
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3101
  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
  3102
proof(rule, rule, rule ccontr)
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3103
  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
  3104
  def x \<equiv> "helper_1 s e"
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3105
  { fix n
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3106
    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
  3107
    proof(induct_tac rule:nat_less_induct)
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3108
      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
  3109
      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
  3110
      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
  3111
      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
  3112
      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
  3113
        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
  3114
      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
  3115
    qed }
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3116
  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
  3117
  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
  3118
  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
  3119
  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
  3120
  show False
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3121
    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
  3122
    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
  3123
    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
  3124
qed
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3125
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3126
subsubsection{* Heine-Borel theorem *}
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3127
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3128
lemma seq_compact_imp_heine_borel:
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3129
  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
  3130
  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
  3131
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
  3132
  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
  3133
  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
  3134
  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
  3135
  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
  3136
    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
  3137
  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
  3138
  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
  3139
    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
  3140
      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
  3141
      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
  3142
               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
  3143
    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
  3144
  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
  3145
    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
  3146
    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
  3147
    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
  3148
    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
  3149
    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
  3150
      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
  3151
    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
  3152
    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
  3153
    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
  3154
      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
  3155
      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
  3156
        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
  3157
      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
  3158
    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
  3159
  qed
50940
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3160
qed
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3161
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3162
lemma compact_eq_seq_compact_metric:
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3163
  "compact (s :: 'a::metric_space set) \<longleftrightarrow> seq_compact s"
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3164
  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
  3165
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3166
lemma compact_def:
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3167
  "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
  3168
   (\<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
  3169
  unfolding compact_eq_seq_compact_metric seq_compact_def by auto
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3170
50944
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  3171
subsubsection {* Complete the chain of compactness variants *}
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  3172
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  3173
lemma compact_eq_bolzano_weierstrass:
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  3174
  fixes s :: "'a::metric_space set"
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  3175
  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
  3176
proof
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  3177
  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
  3178
next
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  3179
  assume ?rhs thus ?lhs
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  3180
    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
  3181
qed
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  3182
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  3183
lemma bolzano_weierstrass_imp_bounded:
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  3184
  "\<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
  3185
  using compact_imp_bounded unfolding compact_eq_bolzano_weierstrass .
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  3186
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3187
text {*
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3188
  A metric space (or topological vector space) is said to have the
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3189
  Heine-Borel property if every closed and bounded subset is compact.
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3190
*}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3191
44207
ea99698c2070 locale-ize some definitions, so perfect_space and heine_borel can inherit from the proper superclasses
huffman
parents: 44170
diff changeset
  3192
class heine_borel = metric_space +
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3193
  assumes bounded_imp_convergent_subsequence:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3194
    "bounded s \<Longrightarrow> \<forall>n. f n \<in> s
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3195
      \<Longrightarrow> \<exists>l r. subseq r \<and> ((f \<circ> r) ---> l) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3196
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  3197
lemma bounded_closed_imp_seq_compact:
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3198
  fixes s::"'a::heine_borel set"
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  3199
  assumes "bounded s" and "closed s" shows "seq_compact s"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  3200
proof (unfold seq_compact_def, clarify)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3201
  fix f :: "nat \<Rightarrow> 'a" assume f: "\<forall>n. f n \<in> s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3202
  obtain l r where r: "subseq r" and l: "((f \<circ> r) ---> l) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3203
    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
  3204
  from f have fr: "\<forall>n. (f \<circ> r) n \<in> s" by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3205
  have "l \<in> s" using `closed s` fr l
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3206
    unfolding closed_sequential_limits by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3207
  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
  3208
    using `l \<in> s` r l by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3209
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3210
50944
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  3211
lemma compact_eq_bounded_closed:
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  3212
  fixes s :: "'a::heine_borel set"
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  3213
  shows "compact s \<longleftrightarrow> bounded s \<and> closed s"  (is "?lhs = ?rhs")
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  3214
proof
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  3215
  assume ?lhs thus ?rhs
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  3216
    using compact_imp_closed compact_imp_bounded by blast
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  3217
next
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  3218
  assume ?rhs thus ?lhs
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  3219
    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
  3220
qed
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  3221
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3222
lemma lim_subseq:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3223
  "subseq r \<Longrightarrow> (s ---> l) sequentially \<Longrightarrow> ((s o r) ---> l) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3224
unfolding tendsto_def eventually_sequentially o_def
50937
d249ef928ae1 removed subseq_bigger (replaced by seq_suble)
hoelzl
parents: 50936
diff changeset
  3225
by (metis seq_suble le_trans)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3226
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3227
lemma num_Axiom: "EX! g. g 0 = e \<and> (\<forall>n. g (Suc n) = f n (g n))"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3228
  unfolding Ex1_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3229
  apply (rule_tac x="nat_rec e f" in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3230
  apply (rule conjI)+
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3231
apply (rule def_nat_rec_0, simp)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3232
apply (rule allI, rule def_nat_rec_Suc, simp)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3233
apply (rule allI, rule impI, rule ext)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3234
apply (erule conjE)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3235
apply (induct_tac x)
36362
06475a1547cb fix lots of looping simp calls and other warnings
huffman
parents: 36360
diff changeset
  3236
apply simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3237
apply (erule_tac x="n" in allE)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3238
apply (simp)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3239
done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3240
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3241
lemma convergent_bounded_increasing: fixes s ::"nat\<Rightarrow>real"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3242
  assumes "incseq s" and "\<forall>n. abs(s n) \<le> b"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3243
  shows "\<exists> l. \<forall>e::real>0. \<exists> N. \<forall>n \<ge> N.  abs(s n - l) < e"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3244
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3245
  have "isUb UNIV (range s) b" using assms(2) and abs_le_D1 unfolding isUb_def and setle_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3246
  then obtain t where t:"isLub UNIV (range s) t" using reals_complete[of "range s" ] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3247
  { fix e::real assume "e>0" and as:"\<forall>N. \<exists>n\<ge>N. \<not> \<bar>s n - t\<bar> < e"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3248
    { fix n::nat
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3249
      obtain N where "N\<ge>n" and n:"\<bar>s N - t\<bar> \<ge> e" using as[THEN spec[where x=n]] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3250
      have "t \<ge> s N" using isLub_isUb[OF t, unfolded isUb_def setle_def] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3251
      with n have "s N \<le> t - e" using `e>0` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3252
      hence "s n \<le> t - e" using assms(1)[unfolded incseq_def, THEN spec[where x=n], THEN spec[where x=N]] using `n\<le>N` by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3253
    hence "isUb UNIV (range s) (t - e)" unfolding isUb_def and setle_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3254
    hence False using isLub_le_isUb[OF t, of "t - e"] and `e>0` by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3255
  thus ?thesis by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3256
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3257
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3258
lemma convergent_bounded_monotone: fixes s::"nat \<Rightarrow> real"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3259
  assumes "\<forall>n. abs(s n) \<le> b" and "monoseq s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3260
  shows "\<exists>l. \<forall>e::real>0. \<exists>N. \<forall>n\<ge>N. abs(s n - l) < e"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3261
  using convergent_bounded_increasing[of s b] assms using convergent_bounded_increasing[of "\<lambda>n. - s n" b]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3262
  unfolding monoseq_def incseq_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3263
  apply auto unfolding minus_add_distrib[THEN sym, unfolded diff_minus[THEN sym]]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3264
  unfolding abs_minus_cancel by(rule_tac x="-l" in exI)auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3265
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  3266
(* TODO: merge this lemma with the ones above *)
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  3267
lemma bounded_increasing_convergent: fixes s::"nat \<Rightarrow> real"
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  3268
  assumes "bounded {s n| n::nat. True}"  "\<forall>n. (s n) \<le>(s(Suc n))"
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  3269
  shows "\<exists>l. (s ---> l) sequentially"
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  3270
proof-
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  3271
  obtain a where a:"\<forall>n. \<bar> (s n)\<bar> \<le>  a" using assms(1)[unfolded bounded_iff] 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
  3272
  { fix m::nat
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  3273
    have "\<And> n. n\<ge>m \<longrightarrow>  (s m) \<le> (s n)"
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  3274
      apply(induct_tac n) apply simp using assms(2) apply(erule_tac x="na" in allE)
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  3275
      apply(case_tac "m \<le> na") unfolding not_less_eq_eq by(auto simp add: not_less_eq_eq)  }
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  3276
  hence "\<forall>m n. m \<le> n \<longrightarrow> (s m) \<le> (s n)" 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
  3277
  then obtain l where "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<bar> (s n) - l\<bar> < e" using convergent_bounded_monotone[OF a]
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  3278
    unfolding monoseq_def by auto
44907
93943da0a010 remove redundant lemma Lim_sequentially in favor of lemma LIMSEQ_def
huffman
parents: 44905
diff changeset
  3279
  thus ?thesis unfolding LIMSEQ_def apply(rule_tac x="l" in exI)
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  3280
    unfolding dist_norm  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
  3281
qed
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  3282
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3283
lemma compact_real_lemma:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3284
  assumes "\<forall>n::nat. abs(s n) \<le> b"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3285
  shows "\<exists>(l::real) r. subseq r \<and> ((s \<circ> r) ---> l) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3286
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3287
  obtain r where r:"subseq r" "monoseq (\<lambda>n. s (r n))"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3288
    using seq_monosub[of s] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3289
  thus ?thesis using convergent_bounded_monotone[of "\<lambda>n. s (r n)" b] and assms
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3290
    unfolding tendsto_iff dist_norm eventually_sequentially by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3291
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3292
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3293
instance real :: heine_borel
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3294
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3295
  fix s :: "real set" and f :: "nat \<Rightarrow> real"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3296
  assume s: "bounded s" and f: "\<forall>n. f n \<in> s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3297
  then obtain b where b: "\<forall>n. abs (f n) \<le> b"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3298
    unfolding bounded_iff by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3299
  obtain l :: real and r :: "nat \<Rightarrow> nat" where
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3300
    r: "subseq r" and l: "((f \<circ> r) ---> l) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3301
    using compact_real_lemma [OF b] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3302
  thus "\<exists>l r. subseq r \<and> ((f \<circ> r) ---> l) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3303
    by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3304
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3305
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3306
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
  3307
  fixes f :: "nat \<Rightarrow> 'a::euclidean_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3308
  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
  3309
  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
  3310
        (\<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
  3311
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
  3312
  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
  3313
  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
  3314
  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
  3315
      (\<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
  3316
  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
  3317
  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
  3318
    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
  3319
      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
  3320
    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
  3321
      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
  3322
      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
  3323
    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
  3324
    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
  3325
      using bounded_imp_convergent_subsequence[OF s' f'] unfolding o_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3326
    def r \<equiv> "r1 \<circ> r2" have r:"subseq r"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3327
      using r1 and r2 unfolding r_def o_def subseq_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3328
    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
  3329
    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
  3330
    { 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
  3331
      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
  3332
      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
  3333
      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
  3334
        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
  3335
      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
  3336
        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
  3337
        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
  3338
    }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3339
    ultimately show ?case by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3340
  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
  3341
qed
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  3342
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  3343
instance euclidean_space \<subseteq> heine_borel
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3344
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
  3345
  fix s :: "'a set" and f :: "nat \<Rightarrow> 'a"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3346
  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
  3347
  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
  3348
    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
  3349
    using compact_lemma [OF s f] by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3350
  { 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
  3351
    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
  3352
    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
  3353
      by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3354
    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
  3355
    { 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
  3356
      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
  3357
        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
  3358
      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
  3359
        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
  3360
      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
  3361
        by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3362
    }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3363
    ultimately have "eventually (\<lambda>n. dist (f (r n)) l < e) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3364
      by (rule eventually_elim1)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3365
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3366
  hence *:"((f \<circ> r) ---> l) sequentially" unfolding o_def tendsto_iff by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3367
  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
  3368
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3369
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3370
lemma bounded_fst: "bounded s \<Longrightarrow> bounded (fst ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3371
unfolding bounded_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3372
apply clarify
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3373
apply (rule_tac x="a" in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3374
apply (rule_tac x="e" in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3375
apply clarsimp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3376
apply (drule (1) bspec)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3377
apply (simp add: dist_Pair_Pair)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3378
apply (erule order_trans [OF real_sqrt_sum_squares_ge1])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3379
done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3380
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3381
lemma bounded_snd: "bounded s \<Longrightarrow> bounded (snd ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3382
unfolding bounded_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3383
apply clarify
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3384
apply (rule_tac x="b" in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3385
apply (rule_tac x="e" in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3386
apply clarsimp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3387
apply (drule (1) bspec)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3388
apply (simp add: dist_Pair_Pair)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3389
apply (erule order_trans [OF real_sqrt_sum_squares_ge2])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3390
done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3391
37678
0040bafffdef "prod" and "sum" replace "*" and "+" respectively
haftmann
parents: 37649
diff changeset
  3392
instance prod :: (heine_borel, heine_borel) heine_borel
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3393
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3394
  fix s :: "('a * 'b) set" and f :: "nat \<Rightarrow> 'a * 'b"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3395
  assume s: "bounded s" and f: "\<forall>n. f n \<in> s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3396
  from s have s1: "bounded (fst ` s)" by (rule bounded_fst)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3397
  from f have f1: "\<forall>n. fst (f n) \<in> fst ` s" by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3398
  obtain l1 r1 where r1: "subseq r1"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3399
    and l1: "((\<lambda>n. fst (f (r1 n))) ---> l1) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3400
    using bounded_imp_convergent_subsequence [OF s1 f1]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3401
    unfolding o_def by fast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3402
  from s have s2: "bounded (snd ` s)" by (rule bounded_snd)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3403
  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
  3404
  obtain l2 r2 where r2: "subseq r2"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3405
    and l2: "((\<lambda>n. snd (f (r1 (r2 n)))) ---> l2) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3406
    using bounded_imp_convergent_subsequence [OF s2 f2]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3407
    unfolding o_def by fast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3408
  have l1': "((\<lambda>n. fst (f (r1 (r2 n)))) ---> l1) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3409
    using lim_subseq [OF r2 l1] unfolding o_def .
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3410
  have l: "((f \<circ> (r1 \<circ> r2)) ---> (l1, l2)) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3411
    using tendsto_Pair [OF l1' l2] unfolding o_def by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3412
  have r: "subseq (r1 \<circ> r2)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3413
    using r1 r2 unfolding subseq_def by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3414
  show "\<exists>l r. subseq r \<and> ((f \<circ> r) ---> l) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3415
    using l r by fast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3416
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3417
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  3418
subsubsection{* Completeness *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3419
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3420
definition
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3421
  complete :: "'a::metric_space set \<Rightarrow> bool" where
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3422
  "complete s \<longleftrightarrow> (\<forall>f. (\<forall>n. f n \<in> s) \<and> Cauchy f
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3423
                      --> (\<exists>l \<in> s. (f ---> l) sequentially))"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3424
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3425
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
  3426
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3427
  { assume ?rhs
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3428
    { fix e::real
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3429
      assume "e>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3430
      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
  3431
        by (erule_tac x="e/2" in allE) auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3432
      { fix n m
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3433
        assume nm:"N \<le> m \<and> N \<le> n"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3434
        hence "dist (s m) (s n) < e" using N
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3435
          using dist_triangle_half_l[of "s m" "s N" "e" "s n"]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3436
          by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3437
      }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3438
      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
  3439
        by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3440
    }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3441
    hence ?lhs
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3442
      unfolding cauchy_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3443
      by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3444
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3445
  thus ?thesis
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3446
    unfolding cauchy_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3447
    using dist_triangle_half_l
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3448
    by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3449
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3450
34104
22758f95e624 re-state lemmas using 'range'
huffman
parents: 33758
diff changeset
  3451
lemma cauchy_imp_bounded: assumes "Cauchy s" shows "bounded (range s)"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3452
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3453
  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
  3454
  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
  3455
  moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3456
  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
  3457
  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
  3458
    unfolding bounded_any_center [where a="s N"] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3459
  ultimately show "?thesis"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3460
    unfolding bounded_any_center [where a="s N"]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3461
    apply(rule_tac x="max a 1" in exI) apply auto
34104
22758f95e624 re-state lemmas using 'range'
huffman
parents: 33758
diff changeset
  3462
    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
  3463
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3464
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  3465
lemma seq_compact_imp_complete: assumes "seq_compact s" shows "complete s"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3466
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3467
  { fix f assume as: "(\<forall>n::nat. f n \<in> s)" "Cauchy f"
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  3468
    from as(1) obtain l r where lr: "l\<in>s" "subseq r" "((f \<circ> r) ---> l) sequentially" using assms unfolding seq_compact_def by blast
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3469
50937
d249ef928ae1 removed subseq_bigger (replaced by seq_suble)
hoelzl
parents: 50936
diff changeset
  3470
    note lr' = seq_suble [OF lr(2)]
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3471
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3472
    { fix e::real assume "e>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3473
      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
44907
93943da0a010 remove redundant lemma Lim_sequentially in favor of lemma LIMSEQ_def
huffman
parents: 44905
diff changeset
  3474
      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
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3475
      { fix n::nat assume n:"n \<ge> max N M"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3476
        have "dist ((f \<circ> r) n) l < e/2" using n M by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3477
        moreover have "r n \<ge> N" using lr'[of n] n by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3478
        hence "dist (f n) ((f \<circ> r) n) < e / 2" using N using n by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3479
        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)  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3480
      hence "\<exists>N. \<forall>n\<ge>N. dist (f n) l < e" by blast  }
44907
93943da0a010 remove redundant lemma Lim_sequentially in favor of lemma LIMSEQ_def
huffman
parents: 44905
diff changeset
  3481
    hence "\<exists>l\<in>s. (f ---> l) sequentially" using `l\<in>s` unfolding LIMSEQ_def by auto  }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3482
  thus ?thesis unfolding complete_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3483
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3484
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3485
instance heine_borel < complete_space
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3486
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3487
  fix f :: "nat \<Rightarrow> 'a" assume "Cauchy f"
34104
22758f95e624 re-state lemmas using 'range'
huffman
parents: 33758
diff changeset
  3488
  hence "bounded (range f)"
22758f95e624 re-state lemmas using 'range'
huffman
parents: 33758
diff changeset
  3489
    by (rule cauchy_imp_bounded)
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  3490
  hence "seq_compact (closure (range f))"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  3491
    using bounded_closed_imp_seq_compact [of "closure (range f)"] by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3492
  hence "complete (closure (range f))"
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  3493
    by (rule seq_compact_imp_complete)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3494
  moreover have "\<forall>n. f n \<in> closure (range f)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3495
    using closure_subset [of "range f"] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3496
  ultimately have "\<exists>l\<in>closure (range f). (f ---> l) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3497
    using `Cauchy f` unfolding complete_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3498
  then show "convergent f"
36660
1cc4ab4b7ff7 make (X ----> L) an abbreviation for (X ---> L) sequentially
huffman
parents: 36659
diff changeset
  3499
    unfolding convergent_def by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3500
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3501
44632
076a45f65e12 simplify/generalize some proofs
huffman
parents: 44628
diff changeset
  3502
instance euclidean_space \<subseteq> banach ..
076a45f65e12 simplify/generalize some proofs
huffman
parents: 44628
diff changeset
  3503
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3504
lemma complete_univ: "complete (UNIV :: 'a::complete_space set)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3505
proof(simp add: complete_def, rule, rule)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3506
  fix f :: "nat \<Rightarrow> 'a" assume "Cauchy f"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3507
  hence "convergent f" by (rule Cauchy_convergent)
36660
1cc4ab4b7ff7 make (X ----> L) an abbreviation for (X ---> L) sequentially
huffman
parents: 36659
diff changeset
  3508
  thus "\<exists>l. f ----> l" unfolding convergent_def .  
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3509
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3510
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3511
lemma complete_imp_closed: assumes "complete s" shows "closed s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3512
proof -
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3513
  { fix x assume "x islimpt s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3514
    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
  3515
      unfolding islimpt_sequential by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3516
    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
  3517
      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
  3518
    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
  3519
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3520
  thus "closed s" unfolding closed_limpt by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3521
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3522
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3523
lemma complete_eq_closed:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3524
  fixes s :: "'a::complete_space set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3525
  shows "complete s \<longleftrightarrow> closed s" (is "?lhs = ?rhs")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3526
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3527
  assume ?lhs thus ?rhs by (rule complete_imp_closed)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3528
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3529
  assume ?rhs
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3530
  { fix f assume as:"\<forall>n::nat. f n \<in> s" "Cauchy f"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3531
    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
  3532
    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
  3533
  thus ?lhs unfolding complete_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3534
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3535
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3536
lemma convergent_eq_cauchy:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3537
  fixes s :: "nat \<Rightarrow> 'a::complete_space"
44632
076a45f65e12 simplify/generalize some proofs
huffman
parents: 44628
diff changeset
  3538
  shows "(\<exists>l. (s ---> l) sequentially) \<longleftrightarrow> Cauchy s"
076a45f65e12 simplify/generalize some proofs
huffman
parents: 44628
diff changeset
  3539
  unfolding Cauchy_convergent_iff convergent_def ..
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3540
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3541
lemma convergent_imp_bounded:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3542
  fixes s :: "nat \<Rightarrow> 'a::metric_space"
44632
076a45f65e12 simplify/generalize some proofs
huffman
parents: 44628
diff changeset
  3543
  shows "(s ---> l) sequentially \<Longrightarrow> bounded (range s)"
50939
ae7cd20ed118 replace convergent_imp_cauchy by LIMSEQ_imp_Cauchy
hoelzl
parents: 50938
diff changeset
  3544
  by (intro cauchy_imp_bounded LIMSEQ_imp_Cauchy)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3545
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3546
lemma nat_approx_posE:
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3547
  fixes e::real
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3548
  assumes "0 < e"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3549
  obtains n::nat where "1 / (Suc n) < e"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3550
proof atomize_elim
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3551
  have " 1 / real (Suc (nat (ceiling (1/e)))) < 1 / (ceiling (1/e))"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3552
    by (rule divide_strict_left_mono) (auto intro!: mult_pos_pos simp: `0 < e`)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3553
  also have "1 / (ceiling (1/e)) \<le> 1 / (1/e)"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3554
    by (rule divide_left_mono) (auto intro!: divide_pos_pos simp: `0 < e`)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3555
  also have "\<dots> = e" by simp
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3556
  finally show  "\<exists>n. 1 / real (Suc n) < e" ..
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3557
qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3558
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3559
lemma compact_eq_totally_bounded:
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  3560
  "compact s \<longleftrightarrow> complete s \<and> (\<forall>e>0. \<exists>k. finite k \<and> s \<subseteq> (\<Union>((\<lambda>x. ball x e) ` k)))"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  3561
proof (safe intro!: seq_compact_imp_complete[unfolded  compact_eq_seq_compact_metric[symmetric]])
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3562
  fix e::real
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3563
  def f \<equiv> "(\<lambda>x::'a. ball x e) ` UNIV"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3564
  assume "0 < e" "compact s"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3565
  hence "(\<forall>t\<in>f. open t) \<and> s \<subseteq> \<Union>f \<longrightarrow> (\<exists>f'\<subseteq>f. finite f' \<and> s \<subseteq> \<Union>f')"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3566
    by (simp add: compact_eq_heine_borel)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3567
  moreover
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3568
  have d0: "\<And>x::'a. dist x x < e" using `0 < e` by simp
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3569
  hence "(\<forall>t\<in>f. open t) \<and> s \<subseteq> \<Union>f" by (auto simp: f_def intro!: d0)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3570
  ultimately have "(\<exists>f'\<subseteq>f. finite f' \<and> s \<subseteq> \<Union>f')" ..
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3571
  then guess K .. note K = this
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3572
  have "\<forall>K'\<in>K. \<exists>k. K' = ball k e" using K by (auto simp: f_def)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3573
  then obtain k where "\<And>K'. K' \<in> K \<Longrightarrow> K' = ball (k K') e" unfolding bchoice_iff by blast
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3574
  thus "\<exists>k. finite k \<and> s \<subseteq> \<Union>(\<lambda>x. ball x e) ` k" using K
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3575
    by (intro exI[where x="k ` K"]) (auto simp: f_def)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3576
next
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3577
  assume assms: "complete s" "\<forall>e>0. \<exists>k. finite k \<and> s \<subseteq> \<Union>(\<lambda>x. ball x e) ` k"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3578
  show "compact s"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3579
  proof cases
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3580
    assume "s = {}" thus "compact s" by (simp add: compact_def)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3581
  next
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3582
    assume "s \<noteq> {}"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3583
    show ?thesis
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3584
      unfolding compact_def
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3585
    proof safe
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3586
      fix f::"nat \<Rightarrow> _" assume "\<forall>n. f n \<in> s" hence f: "\<And>n. f n \<in> s" by simp
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3587
      from assms have "\<forall>e. \<exists>k. e>0 \<longrightarrow> finite k \<and> s \<subseteq> (\<Union>((\<lambda>x. ball x e) ` k))" by simp
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3588
      then obtain K where
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3589
        K: "\<And>e. e > 0 \<Longrightarrow> finite (K e) \<and> s \<subseteq> (\<Union>((\<lambda>x. ball x e) ` (K e)))"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3590
        unfolding choice_iff by blast
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3591
      {
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3592
        fix e::real and f' have f': "\<And>n::nat. (f o f') n \<in> s" using f by auto
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3593
        assume "e > 0"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3594
        from K[OF this] have K: "finite (K e)" "s \<subseteq> (\<Union>((\<lambda>x. ball x e) ` (K e)))"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3595
          by simp_all
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3596
        have "\<exists>k\<in>(K e). \<exists>r. subseq r \<and> (\<forall>i. (f o f' o r) i \<in> ball k e)"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3597
        proof (rule ccontr)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3598
          from K have "finite (K e)" "K e \<noteq> {}" "s \<subseteq> (\<Union>((\<lambda>x. ball x e) ` (K e)))"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3599
            using `s \<noteq> {}`
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3600
            by auto
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3601
          moreover
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3602
          assume "\<not> (\<exists>k\<in>K e. \<exists>r. subseq r \<and> (\<forall>i. (f \<circ> f' o r) i \<in> ball k e))"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3603
          hence "\<And>r k. k \<in> K e \<Longrightarrow> subseq r \<Longrightarrow> (\<exists>i. (f o f' o r) i \<notin> ball k e)" by simp
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3604
          ultimately
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3605
          show False using f'
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3606
          proof (induct arbitrary: s f f' rule: finite_ne_induct)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3607
            case (singleton x)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3608
            have "\<exists>i. (f \<circ> f' o id) i \<notin> ball x e" by (rule singleton) (auto simp: subseq_def)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3609
            thus ?case using singleton by (auto simp: ball_def)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3610
          next
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3611
            case (insert x A)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3612
            show ?case
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3613
            proof cases
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3614
              have inf_ms: "infinite ((f o f') -` s)" using insert by (simp add: vimage_def)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3615
              have "infinite ((f o f') -` \<Union>((\<lambda>x. ball x e) ` (insert x A)))"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3616
                using insert by (intro infinite_super[OF _ inf_ms]) auto
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3617
              also have "((f o f') -` \<Union>((\<lambda>x. ball x e) ` (insert x A))) =
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3618
                {m. (f o f') m \<in> ball x e} \<union> {m. (f o f') m \<in> \<Union>((\<lambda>x. ball x e) ` A)}" by auto
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3619
              finally have "infinite \<dots>" .
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3620
              moreover assume "finite {m. (f o f') m \<in> ball x e}"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3621
              ultimately have inf: "infinite {m. (f o f') m \<in> \<Union>((\<lambda>x. ball x e) ` A)}" by blast
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3622
              hence "A \<noteq> {}" by auto then obtain k where "k \<in> A" by auto
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3623
              def r \<equiv> "enumerate {m. (f o f') m \<in> \<Union>((\<lambda>x. ball x e) ` A)}"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3624
              have r_mono: "\<And>n m. n < m \<Longrightarrow> r n < r m"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3625
                using enumerate_mono[OF _ inf] by (simp add: r_def)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3626
              hence "subseq r" by (simp add: subseq_def)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3627
              have r_in_set: "\<And>n. r n \<in> {m. (f o f') m \<in> \<Union>((\<lambda>x. ball x e) ` A)}"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3628
                using enumerate_in_set[OF inf] by (simp add: r_def)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3629
              show False
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3630
              proof (rule insert)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3631
                show "\<Union>(\<lambda>x. ball x e) ` A \<subseteq> \<Union>(\<lambda>x. ball x e) ` A" by simp
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3632
                fix k s assume "k \<in> A" "subseq s"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3633
                thus "\<exists>i. (f o f' o r o s) i \<notin> ball k e" using `subseq r`
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3634
                  by (subst (2) o_assoc[symmetric]) (intro insert(6) subseq_o, simp_all)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3635
              next
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3636
                fix n show "(f \<circ> f' o r) n \<in> \<Union>(\<lambda>x. ball x e) ` A" using r_in_set by auto
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3637
              qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3638
            next
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3639
              assume inf: "infinite {m. (f o f') m \<in> ball x e}"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3640
              def r \<equiv> "enumerate {m. (f o f') m \<in> ball x e}"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3641
              have r_mono: "\<And>n m. n < m \<Longrightarrow> r n < r m"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3642
                using enumerate_mono[OF _ inf] by (simp add: r_def)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3643
              hence "subseq r" by (simp add: subseq_def)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3644
              from insert(6)[OF insertI1 this] obtain i where "(f o f') (r i) \<notin> ball x e" by auto
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3645
              moreover
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3646
              have r_in_set: "\<And>n. r n \<in> {m. (f o f') m \<in> ball x e}"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3647
                using enumerate_in_set[OF inf] by (simp add: r_def)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3648
              hence "(f o f') (r i) \<in> ball x e" by simp
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3649
              ultimately show False by simp
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3650
            qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3651
          qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3652
        qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3653
      }
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3654
      hence ex: "\<forall>f'. \<forall>e > 0. (\<exists>k\<in>K e. \<exists>r. subseq r \<and> (\<forall>i. (f o f' \<circ> r) i \<in> ball k e))" by simp
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3655
      let ?e = "\<lambda>n. 1 / real (Suc n)"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3656
      let ?P = "\<lambda>n s. \<exists>k\<in>K (?e n). (\<forall>i. (f o s) i \<in> ball k (?e n))"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3657
      interpret subseqs ?P using ex by unfold_locales force
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3658
      from `complete s` have limI: "\<And>f. (\<And>n. f n \<in> s) \<Longrightarrow> Cauchy f \<Longrightarrow> (\<exists>l\<in>s. f ----> l)"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3659
        by (simp add: complete_def)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3660
      have "\<exists>l\<in>s. (f o diagseq) ----> l"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3661
      proof (intro limI metric_CauchyI)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3662
        fix e::real assume "0 < e" hence "0 < e / 2" by auto
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3663
        from nat_approx_posE[OF this] guess n . note n = this
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3664
        show "\<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. dist ((f \<circ> diagseq) m) ((f \<circ> diagseq) n) < e"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3665
        proof (rule exI[where x="Suc n"], safe)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3666
          fix m mm assume "Suc n \<le> m" "Suc n \<le> mm"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3667
          let ?e = "1 / real (Suc n)"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3668
          from reducer_reduces[of n] obtain k where
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3669
            "k\<in>K ?e"  "\<And>i. (f o seqseq (Suc n)) i \<in> ball k ?e"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3670
            unfolding seqseq_reducer by auto
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3671
          moreover
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3672
          note diagseq_sub[OF `Suc n \<le> m`] diagseq_sub[OF `Suc n \<le> mm`]
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3673
          ultimately have "{(f o diagseq) m, (f o diagseq) mm} \<subseteq> ball k ?e" by auto
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3674
          also have "\<dots> \<subseteq> ball k (e / 2)" using n by (intro subset_ball) simp
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3675
          finally
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3676
          have "dist k ((f \<circ> diagseq) m) + dist k ((f \<circ> diagseq) mm) < e / 2 + e /2"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3677
            by (intro add_strict_mono) auto
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3678
          hence "dist ((f \<circ> diagseq) m) k + dist ((f \<circ> diagseq) mm) k < e"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3679
            by (simp add: dist_commute)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3680
          moreover have "dist ((f \<circ> diagseq) m) ((f \<circ> diagseq) mm) \<le>
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3681
            dist ((f \<circ> diagseq) m) k + dist ((f \<circ> diagseq) mm) k"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3682
            by (rule dist_triangle2)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3683
          ultimately show "dist ((f \<circ> diagseq) m) ((f \<circ> diagseq) mm) < e"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3684
            by simp
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3685
        qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3686
      next
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3687
        fix n show "(f o diagseq) n \<in> s" using f by simp
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3688
      qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3689
      thus "\<exists>l\<in>s. \<exists>r. subseq r \<and> (f \<circ> r) ----> l" using subseq_diagseq by auto
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3690
    qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3691
  qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3692
qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3693
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3694
lemma compact_cball[simp]:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3695
  fixes x :: "'a::heine_borel"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3696
  shows "compact(cball x e)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3697
  using compact_eq_bounded_closed bounded_cball closed_cball
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3698
  by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3699
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3700
lemma compact_frontier_bounded[intro]:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3701
  fixes s :: "'a::heine_borel set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3702
  shows "bounded s ==> compact(frontier s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3703
  unfolding frontier_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3704
  using compact_eq_bounded_closed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3705
  by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3706
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3707
lemma compact_frontier[intro]:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3708
  fixes s :: "'a::heine_borel set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3709
  shows "compact s ==> compact (frontier s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3710
  using compact_eq_bounded_closed compact_frontier_bounded
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3711
  by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3712
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3713
lemma frontier_subset_compact:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3714
  fixes s :: "'a::heine_borel set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3715
  shows "compact s ==> frontier s \<subseteq> s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3716
  using frontier_subset_closed compact_eq_bounded_closed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3717
  by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3718
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  3719
subsection {* Bounded closed nest property (proof does not use Heine-Borel) *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3720
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3721
lemma bounded_closed_nest:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3722
  assumes "\<forall>n. closed(s n)" "\<forall>n. (s n \<noteq> {})"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3723
  "(\<forall>m n. m \<le> n --> s n \<subseteq> s m)"  "bounded(s 0)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3724
  shows "\<exists>a::'a::heine_borel. \<forall>n::nat. a \<in> s(n)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3725
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3726
  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
  3727
  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
  3728
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3729
  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
  3730
    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
  3731
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3732
  { fix n::nat
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3733
    { fix e::real assume "e>0"
44907
93943da0a010 remove redundant lemma Lim_sequentially in favor of lemma LIMSEQ_def
huffman
parents: 44905
diff changeset
  3734
      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
  3735
      hence "dist ((x \<circ> r) (max N n)) l < e" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3736
      moreover
50937
d249ef928ae1 removed subseq_bigger (replaced by seq_suble)
hoelzl
parents: 50936
diff changeset
  3737
      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
  3738
      hence "(x \<circ> r) (max N n) \<in> s n"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3739
        using x apply(erule_tac x=n in allE)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3740
        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
  3741
        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
  3742
      ultimately have "\<exists>y\<in>s n. dist y l < e" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3743
    }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3744
    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
  3745
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3746
  thus ?thesis by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3747
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3748
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  3749
text {* Decreasing case does not even need compactness, just completeness. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3750
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3751
lemma decreasing_closed_nest:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3752
  assumes "\<forall>n. closed(s n)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3753
          "\<forall>n. (s n \<noteq> {})"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3754
          "\<forall>m n. m \<le> n --> s n \<subseteq> s m"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3755
          "\<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
  3756
  shows "\<exists>a::'a::complete_space. \<forall>n::nat. a \<in> s n"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3757
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3758
  have "\<forall>n. \<exists> x. x\<in>s n" using assms(2) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3759
  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
  3760
  then obtain t where t: "\<forall>n. t n \<in> s n" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3761
  { fix e::real assume "e>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3762
    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
  3763
    { fix m n ::nat assume "N \<le> m \<and> N \<le> n"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3764
      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
  3765
      hence "dist (t m) (t n) < e" using N by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3766
    }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3767
    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
  3768
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3769
  hence  "Cauchy t" unfolding cauchy_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3770
  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
  3771
  { fix n::nat
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3772
    { fix e::real assume "e>0"
44907
93943da0a010 remove redundant lemma Lim_sequentially in favor of lemma LIMSEQ_def
huffman
parents: 44905
diff changeset
  3773
      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
  3774
      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
  3775
      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
  3776
    }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3777
    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
  3778
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3779
  then show ?thesis by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3780
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3781
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  3782
text {* Strengthen it to the intersection actually being a singleton. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3783
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3784
lemma decreasing_closed_nest_sing:
44632
076a45f65e12 simplify/generalize some proofs
huffman
parents: 44628
diff changeset
  3785
  fixes s :: "nat \<Rightarrow> 'a::complete_space set"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3786
  assumes "\<forall>n. closed(s n)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3787
          "\<forall>n. s n \<noteq> {}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3788
          "\<forall>m n. m \<le> n --> s n \<subseteq> s m"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3789
          "\<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
  3790
  shows "\<exists>a. \<Inter>(range s) = {a}"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3791
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3792
  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
  3793
  { fix b assume b:"b \<in> \<Inter>(range s)"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3794
    { fix e::real assume "e>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3795
      hence "dist a b < e" using assms(4 )using b using a by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3796
    }
36778
739a9379e29b avoid using real-specific versions of generic lemmas
huffman
parents: 36669
diff changeset
  3797
    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
  3798
  }
34104
22758f95e624 re-state lemmas using 'range'
huffman
parents: 33758
diff changeset
  3799
  with a have "\<Inter>(range s) = {a}" unfolding image_def by auto
22758f95e624 re-state lemmas using 'range'
huffman
parents: 33758
diff changeset
  3800
  thus ?thesis ..
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3801
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3802
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3803
text{* Cauchy-type criteria for uniform convergence. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3804
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3805
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
  3806
 "(\<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
  3807
  (\<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
  3808
proof(rule)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3809
  assume ?lhs
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3810
  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
  3811
  { fix e::real assume "e>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3812
    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
  3813
    { 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
  3814
      hence "dist (s m x) (s n x) < e"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3815
        using N[THEN spec[where x=m], THEN spec[where x=x]]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3816
        using N[THEN spec[where x=n], THEN spec[where x=x]]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3817
        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
  3818
    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
  3819
  thus ?rhs by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3820
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3821
  assume ?rhs
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3822
  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
  3823
  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
  3824
    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
  3825
  { fix e::real assume "e>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3826
    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
  3827
      using `?rhs`[THEN spec[where x="e/2"]] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3828
    { fix x assume "P x"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3829
      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
  3830
        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
  3831
      fix n::nat assume "n\<ge>N"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3832
      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
  3833
        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
  3834
    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
  3835
  thus ?lhs by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3836
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3837
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3838
lemma uniformly_cauchy_imp_uniformly_convergent:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3839
  fixes s :: "nat \<Rightarrow> 'a \<Rightarrow> 'b::heine_borel"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3840
  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
  3841
          "\<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
  3842
  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
  3843
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3844
  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
  3845
    using assms(1) unfolding uniformly_convergent_eq_cauchy[THEN sym] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3846
  moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3847
  { fix x assume "P x"
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41969
diff changeset
  3848
    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
  3849
      using l and assms(2) unfolding LIMSEQ_def by blast  }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3850
  ultimately show ?thesis by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3851
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3852
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  3853
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  3854
subsection {* Continuity *}
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  3855
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  3856
text {* Define continuity over a net to take in restrictions of the set. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3857
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3858
definition
44081
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44076
diff changeset
  3859
  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
  3860
  where "continuous net f \<longleftrightarrow> (f ---> f(netlimit net)) net"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3861
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3862
lemma continuous_trivial_limit:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3863
 "trivial_limit net ==> continuous net f"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3864
  unfolding continuous_def tendsto_def trivial_limit_eq by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3865
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3866
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
  3867
  unfolding continuous_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3868
  unfolding tendsto_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3869
  using netlimit_within[of x s]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3870
  by (cases "trivial_limit (at x within s)") (auto simp add: trivial_limit_eventually)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3871
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3872
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
  3873
  using continuous_within [of x UNIV f] by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3874
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3875
lemma continuous_at_within:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3876
  assumes "continuous (at x) f"  shows "continuous (at x within s) f"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3877
  using assms unfolding continuous_at continuous_within
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3878
  by (rule Lim_at_within)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3879
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3880
text{* Derive the epsilon-delta forms, which we often use as "definitions" *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3881
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3882
lemma continuous_within_eps_delta:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3883
  "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
  3884
  unfolding continuous_within and Lim_within
44584
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  3885
  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
  3886
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3887
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
  3888
                           \<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
  3889
  using continuous_within_eps_delta [of x UNIV f] by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3890
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3891
text{* Versions in terms of open balls. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3892
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3893
lemma continuous_within_ball:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3894
 "continuous (at x within s) f \<longleftrightarrow> (\<forall>e>0. \<exists>d>0.
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3895
                            f ` (ball x d \<inter> s) \<subseteq> ball (f x) e)" (is "?lhs = ?rhs")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3896
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3897
  assume ?lhs
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3898
  { fix e::real assume "e>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3899
    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
  3900
      using `?lhs`[unfolded continuous_within Lim_within] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3901
    { fix y assume "y\<in>f ` (ball x d \<inter> s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3902
      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
  3903
        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
  3904
    }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3905
    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
  3906
  thus ?rhs by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3907
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3908
  assume ?rhs thus ?lhs unfolding continuous_within Lim_within ball_def subset_eq
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3909
    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
  3910
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3911
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3912
lemma continuous_at_ball:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3913
  "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
  3914
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3915
  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
  3916
    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
  3917
    unfolding dist_nz[THEN sym] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3918
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3919
  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
  3920
    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
  3921
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3922
36440
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3923
text{* Define setwise continuity in terms of limits within the set. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3924
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3925
definition
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  3926
  continuous_on ::
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  3927
    "'a set \<Rightarrow> ('a::topological_space \<Rightarrow> 'b::topological_space) \<Rightarrow> bool"
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  3928
where
36440
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3929
  "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
  3930
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3931
lemma continuous_on_topological:
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  3932
  "continuous_on s f \<longleftrightarrow>
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  3933
    (\<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
  3934
      (\<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
  3935
unfolding continuous_on_def tendsto_def
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3936
unfolding Limits.eventually_within eventually_at_topological
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3937
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
  3938
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  3939
lemma continuous_on_iff:
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  3940
  "continuous_on s f \<longleftrightarrow>
36440
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3941
    (\<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
  3942
unfolding continuous_on_def Lim_within
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3943
apply (intro ball_cong [OF refl] all_cong ex_cong)
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3944
apply (rename_tac y, case_tac "y = x", simp)
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3945
apply (simp add: dist_nz)
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  3946
done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3947
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3948
definition
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3949
  uniformly_continuous_on ::
36440
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3950
    "'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
  3951
where
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3952
  "uniformly_continuous_on s f \<longleftrightarrow>
36440
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3953
    (\<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
  3954
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3955
text{* Some simple consequential lemmas. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3956
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3957
lemma uniformly_continuous_imp_continuous:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3958
 " uniformly_continuous_on s f ==> continuous_on s f"
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  3959
  unfolding uniformly_continuous_on_def continuous_on_iff by blast
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3960
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3961
lemma continuous_at_imp_continuous_within:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3962
 "continuous (at x) f ==> continuous (at x within s) f"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3963
  unfolding continuous_within continuous_at using Lim_at_within by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3964
36440
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3965
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
  3966
unfolding tendsto_def by (simp add: trivial_limit_eq)
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3967
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  3968
lemma continuous_at_imp_continuous_on:
36440
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3969
  assumes "\<forall>x\<in>s. continuous (at x) f"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3970
  shows "continuous_on s f"
36440
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3971
unfolding continuous_on_def
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3972
proof
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3973
  fix x assume "x \<in> s"
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3974
  with assms have *: "(f ---> f (netlimit (at x))) (at x)"
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3975
    unfolding continuous_def by simp
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3976
  have "(f ---> f x) (at x)"
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3977
  proof (cases "trivial_limit (at x)")
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3978
    case True thus ?thesis
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3979
      by (rule Lim_trivial_limit)
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3980
  next
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3981
    case False
36667
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  3982
    hence 1: "netlimit (at x) = x"
45031
9583f2b56f85 add lemmas within_empty and tendsto_bot;
huffman
parents: 44909
diff changeset
  3983
      using netlimit_within [of x UNIV] by simp
36440
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3984
    with * show ?thesis by simp
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3985
  qed
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3986
  thus "(f ---> f x) (at x within s)"
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3987
    by (rule Lim_at_within)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3988
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3989
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3990
lemma continuous_on_eq_continuous_within:
36440
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3991
  "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
  3992
unfolding continuous_on_def continuous_def
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3993
apply (rule ball_cong [OF refl])
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3994
apply (case_tac "trivial_limit (at x within s)")
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3995
apply (simp add: Lim_trivial_limit)
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3996
apply (simp add: netlimit_within)
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3997
done
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3998
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3999
lemmas continuous_on = continuous_on_def -- "legacy theorem name"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4000
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4001
lemma continuous_on_eq_continuous_at:
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  4002
  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
  4003
  by (auto simp add: continuous_on continuous_at Lim_within_open)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4004
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4005
lemma continuous_within_subset:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4006
 "continuous (at x within s) f \<Longrightarrow> t \<subseteq> s
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4007
             ==> continuous (at x within t) f"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4008
  unfolding continuous_within by(metis Lim_within_subset)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4009
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4010
lemma continuous_on_subset:
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  4011
  shows "continuous_on s f \<Longrightarrow> t \<subseteq> s ==> continuous_on t f"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4012
  unfolding continuous_on by (metis subset_eq Lim_within_subset)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4013
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4014
lemma continuous_on_interior:
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  4015
  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
  4016
  by (erule interiorE, drule (1) continuous_on_subset,
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  4017
    simp add: continuous_on_eq_continuous_at)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4018
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4019
lemma continuous_on_eq:
36440
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  4020
  "(\<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
  4021
  unfolding continuous_on_def tendsto_def Limits.eventually_within
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  4022
  by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4023
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  4024
text {* Characterization of various kinds of continuity in terms of sequences. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4025
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4026
lemma continuous_within_sequentially:
44533
7abe4a59f75d generalize and simplify proof of continuous_within_sequentially
huffman
parents: 44531
diff changeset
  4027
  fixes f :: "'a::metric_space \<Rightarrow> 'b::topological_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4028
  shows "continuous (at a within s) f \<longleftrightarrow>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4029
                (\<forall>x. (\<forall>n::nat. x n \<in> s) \<and> (x ---> a) sequentially
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4030
                     --> ((f o x) ---> f a) sequentially)" (is "?lhs = ?rhs")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4031
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4032
  assume ?lhs
44533
7abe4a59f75d generalize and simplify proof of continuous_within_sequentially
huffman
parents: 44531
diff changeset
  4033
  { 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
  4034
    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
  4035
    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
  4036
      unfolding continuous_within tendsto_def eventually_within by auto
7abe4a59f75d generalize and simplify proof of continuous_within_sequentially
huffman
parents: 44531
diff changeset
  4037
    have "eventually (\<lambda>n. dist (x n) a < d) sequentially"
7abe4a59f75d generalize and simplify proof of continuous_within_sequentially
huffman
parents: 44531
diff changeset
  4038
      using x(2) `d>0` by simp
7abe4a59f75d generalize and simplify proof of continuous_within_sequentially
huffman
parents: 44531
diff changeset
  4039
    hence "eventually (\<lambda>n. (f \<circ> x) n \<in> T) sequentially"
46887
cb891d9a23c1 use eventually_elim method
noschinl
parents: 45776
diff changeset
  4040
    proof eventually_elim
cb891d9a23c1 use eventually_elim method
noschinl
parents: 45776
diff changeset
  4041
      case (elim n) thus ?case
44533
7abe4a59f75d generalize and simplify proof of continuous_within_sequentially
huffman
parents: 44531
diff changeset
  4042
        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
  4043
    qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4044
  }
44533
7abe4a59f75d generalize and simplify proof of continuous_within_sequentially
huffman
parents: 44531
diff changeset
  4045
  thus ?rhs unfolding tendsto_iff unfolding tendsto_def by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4046
next
44533
7abe4a59f75d generalize and simplify proof of continuous_within_sequentially
huffman
parents: 44531
diff changeset
  4047
  assume ?rhs thus ?lhs
7abe4a59f75d generalize and simplify proof of continuous_within_sequentially
huffman
parents: 44531
diff changeset
  4048
    unfolding continuous_within tendsto_def [where l="f a"]
7abe4a59f75d generalize and simplify proof of continuous_within_sequentially
huffman
parents: 44531
diff changeset
  4049
    by (simp add: sequentially_imp_eventually_within)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4050
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4051
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4052
lemma continuous_at_sequentially:
44533
7abe4a59f75d generalize and simplify proof of continuous_within_sequentially
huffman
parents: 44531
diff changeset
  4053
  fixes f :: "'a::metric_space \<Rightarrow> 'b::topological_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4054
  shows "continuous (at a) f \<longleftrightarrow> (\<forall>x. (x ---> a) sequentially
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4055
                  --> ((f o x) ---> f a) sequentially)"
45031
9583f2b56f85 add lemmas within_empty and tendsto_bot;
huffman
parents: 44909
diff changeset
  4056
  using continuous_within_sequentially[of a UNIV f] by simp
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_on_sequentially:
44533
7abe4a59f75d generalize and simplify proof of continuous_within_sequentially
huffman
parents: 44531
diff changeset
  4059
  fixes f :: "'a::metric_space \<Rightarrow> 'b::topological_space"
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  4060
  shows "continuous_on s f \<longleftrightarrow>
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  4061
    (\<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
  4062
                    --> ((f o x) ---> f(a)) sequentially)" (is "?lhs = ?rhs")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4063
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4064
  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
  4065
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4066
  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
  4067
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4068
44648
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4069
lemma uniformly_continuous_on_sequentially:
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4070
  "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
  4071
                    ((\<lambda>n. dist (x n) (y n)) ---> 0) sequentially
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4072
                    \<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
  4073
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4074
  assume ?lhs
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4075
  { 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
  4076
    { fix e::real assume "e>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4077
      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
  4078
        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
  4079
      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
  4080
      { fix n assume "n\<ge>N"
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4081
        hence "dist (f (x n)) (f (y n)) < e"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4082
          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
  4083
          unfolding dist_commute by simp  }
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4084
      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
  4085
    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
  4086
  thus ?rhs by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4087
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4088
  assume ?rhs
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4089
  { assume "\<not> ?lhs"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4090
    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
  4091
    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
  4092
      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
  4093
      by (auto simp add: dist_commute)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4094
    def x \<equiv> "\<lambda>n::nat. fst (fa (inverse (real n + 1)))"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4095
    def y \<equiv> "\<lambda>n::nat. snd (fa (inverse (real n + 1)))"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4096
    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
  4097
      unfolding x_def and y_def using fa by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4098
    { fix e::real assume "e>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4099
      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
  4100
      { fix n::nat assume "n\<ge>N"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4101
        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
  4102
        also have "\<dots> < e" using N by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4103
        finally have "inverse (real n + 1) < e" by auto
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4104
        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
  4105
      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
  4106
    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
  4107
    hence False using fxy and `e>0` by auto  }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4108
  thus ?lhs unfolding uniformly_continuous_on_def by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4109
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4110
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4111
text{* The usual transformation theorems. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4112
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4113
lemma continuous_transform_within:
36667
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  4114
  fixes f g :: "'a::metric_space \<Rightarrow> 'b::topological_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4115
  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
  4116
          "continuous (at x within s) f"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4117
  shows "continuous (at x within s) g"
36667
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  4118
unfolding continuous_within
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  4119
proof (rule Lim_transform_within)
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  4120
  show "0 < d" by fact
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  4121
  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
  4122
    using assms(3) by auto
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  4123
  have "f x = g x"
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  4124
    using assms(1,2,3) by auto
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  4125
  thus "(f ---> g x) (at x within s)"
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  4126
    using assms(4) unfolding continuous_within by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4127
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4128
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4129
lemma continuous_transform_at:
36667
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  4130
  fixes f g :: "'a::metric_space \<Rightarrow> 'b::topological_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4131
  assumes "0 < d" "\<forall>x'. dist x' x < d --> f x' = g x'"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4132
          "continuous (at x) f"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4133
  shows "continuous (at x) g"
45031
9583f2b56f85 add lemmas within_empty and tendsto_bot;
huffman
parents: 44909
diff changeset
  4134
  using continuous_transform_within [of d x UNIV f g] assms by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4135
44648
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4136
subsubsection {* Structural rules for pointwise continuity *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4137
44647
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4138
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
  4139
  unfolding continuous_within by (rule tendsto_ident_at_within)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4140
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4141
lemma continuous_at_id: "continuous (at a) (\<lambda>x. x)"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4142
  unfolding continuous_at by (rule tendsto_ident_at)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4143
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4144
lemma continuous_const: "continuous F (\<lambda>x. c)"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4145
  unfolding continuous_def by (rule tendsto_const)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4146
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4147
lemma continuous_dist:
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4148
  assumes "continuous F f" and "continuous F g"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4149
  shows "continuous F (\<lambda>x. dist (f x) (g x))"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4150
  using assms unfolding continuous_def by (rule tendsto_dist)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4151
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  4152
lemma continuous_infdist:
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  4153
  assumes "continuous F f"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  4154
  shows "continuous F (\<lambda>x. infdist (f x) A)"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  4155
  using assms unfolding continuous_def by (rule tendsto_infdist)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  4156
44647
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4157
lemma continuous_norm:
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4158
  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
  4159
  unfolding continuous_def by (rule tendsto_norm)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4160
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4161
lemma continuous_infnorm:
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4162
  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
  4163
  unfolding continuous_def by (rule tendsto_infnorm)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4164
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4165
lemma continuous_add:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4166
  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
  4167
  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
  4168
  unfolding continuous_def by (rule tendsto_add)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4169
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4170
lemma continuous_minus:
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4171
  fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4172
  shows "continuous F f \<Longrightarrow> continuous F (\<lambda>x. - f x)"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4173
  unfolding continuous_def by (rule tendsto_minus)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4174
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4175
lemma continuous_diff:
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4176
  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
  4177
  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
  4178
  unfolding continuous_def by (rule tendsto_diff)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4179
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4180
lemma continuous_scaleR:
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4181
  fixes g :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4182
  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
  4183
  unfolding continuous_def by (rule tendsto_scaleR)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4184
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4185
lemma continuous_mult:
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4186
  fixes f g :: "'a::t2_space \<Rightarrow> 'b::real_normed_algebra"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4187
  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
  4188
  unfolding continuous_def by (rule tendsto_mult)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4189
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4190
lemma continuous_inner:
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4191
  assumes "continuous F f" and "continuous F g"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4192
  shows "continuous F (\<lambda>x. inner (f x) (g x))"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4193
  using assms unfolding continuous_def by (rule tendsto_inner)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4194
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4195
lemma continuous_inverse:
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4196
  fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_div_algebra"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4197
  assumes "continuous F f" and "f (netlimit F) \<noteq> 0"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4198
  shows "continuous F (\<lambda>x. inverse (f x))"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4199
  using assms unfolding continuous_def by (rule tendsto_inverse)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4200
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4201
lemma continuous_at_within_inverse:
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4202
  fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_div_algebra"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4203
  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
  4204
  shows "continuous (at a within s) (\<lambda>x. inverse (f x))"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4205
  using assms unfolding continuous_within by (rule tendsto_inverse)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4206
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4207
lemma continuous_at_inverse:
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4208
  fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_div_algebra"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4209
  assumes "continuous (at a) f" and "f a \<noteq> 0"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4210
  shows "continuous (at a) (\<lambda>x. inverse (f x))"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4211
  using assms unfolding continuous_at by (rule tendsto_inverse)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4212
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4213
lemmas continuous_intros = continuous_at_id continuous_within_id
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4214
  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
  4215
  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
  4216
  continuous_inner continuous_at_inverse continuous_at_within_inverse
34964
4e8be3c04d37 Replaced vec1 and dest_vec1 by abbreviation.
hoelzl
parents: 34291
diff changeset
  4217
44648
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4218
subsubsection {* Structural rules for setwise continuity *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4219
44647
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4220
lemma continuous_on_id: "continuous_on s (\<lambda>x. x)"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4221
  unfolding continuous_on_def by (fast intro: tendsto_ident_at_within)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4222
44531
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4223
lemma continuous_on_const: "continuous_on s (\<lambda>x. c)"
44125
230a8665c919 mark some redundant theorems as legacy
huffman
parents: 44122
diff changeset
  4224
  unfolding continuous_on_def by (auto intro: tendsto_intros)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4225
44647
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4226
lemma continuous_on_norm:
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4227
  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
  4228
  unfolding continuous_on_def by (fast intro: tendsto_norm)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4229
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4230
lemma continuous_on_infnorm:
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4231
  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
  4232
  unfolding continuous_on by (fast intro: tendsto_infnorm)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4233
44531
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4234
lemma continuous_on_minus:
36440
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  4235
  fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4236
  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
  4237
  unfolding continuous_on_def by (auto intro: tendsto_intros)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4238
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4239
lemma continuous_on_add:
36440
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  4240
  fixes f g :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4241
  shows "continuous_on s f \<Longrightarrow> continuous_on s g
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4242
           \<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
  4243
  unfolding continuous_on_def by (auto intro: tendsto_intros)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4244
44531
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4245
lemma continuous_on_diff:
36440
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  4246
  fixes f g :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4247
  shows "continuous_on s f \<Longrightarrow> continuous_on s g
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4248
           \<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
  4249
  unfolding continuous_on_def by (auto intro: tendsto_intros)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4250
44531
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4251
lemma (in bounded_linear) continuous_on:
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4252
  "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
  4253
  unfolding continuous_on_def by (fast intro: tendsto)
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4254
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4255
lemma (in bounded_bilinear) continuous_on:
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4256
  "\<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
  4257
  unfolding continuous_on_def by (fast intro: tendsto)
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4258
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4259
lemma continuous_on_scaleR:
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4260
  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
  4261
  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
  4262
  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
  4263
  using bounded_bilinear_scaleR assms
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4264
  by (rule bounded_bilinear.continuous_on)
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4265
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4266
lemma continuous_on_mult:
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4267
  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
  4268
  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
  4269
  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
  4270
  using bounded_bilinear_mult assms
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4271
  by (rule bounded_bilinear.continuous_on)
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4272
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4273
lemma continuous_on_inner:
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4274
  fixes g :: "'a::topological_space \<Rightarrow> 'b::real_inner"
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4275
  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
  4276
  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
  4277
  using bounded_bilinear_inner assms
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4278
  by (rule bounded_bilinear.continuous_on)
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4279
44647
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4280
lemma continuous_on_inverse:
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4281
  fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_div_algebra"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4282
  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
  4283
  shows "continuous_on s (\<lambda>x. inverse (f x))"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4284
  using assms unfolding continuous_on by (fast intro: tendsto_inverse)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4285
44648
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4286
subsubsection {* Structural rules for uniform continuity *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4287
44647
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4288
lemma uniformly_continuous_on_id:
44648
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4289
  shows "uniformly_continuous_on s (\<lambda>x. x)"
44647
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4290
  unfolding uniformly_continuous_on_def by auto
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4291
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4292
lemma uniformly_continuous_on_const:
44648
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4293
  shows "uniformly_continuous_on s (\<lambda>x. c)"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4294
  unfolding uniformly_continuous_on_def by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4295
44648
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4296
lemma uniformly_continuous_on_dist:
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4297
  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
  4298
  assumes "uniformly_continuous_on s f"
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4299
  assumes "uniformly_continuous_on s g"
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4300
  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
  4301
proof -
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4302
  { 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
  4303
      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
  4304
      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
  4305
      by arith
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4306
  } note le = this
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4307
  { fix x y
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4308
    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
  4309
    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
  4310
    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
  4311
      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
  4312
        simp add: le)
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4313
  }
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4314
  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
  4315
    unfolding dist_real_def by simp
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4316
qed
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4317
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4318
lemma uniformly_continuous_on_norm:
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4319
  assumes "uniformly_continuous_on s f"
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4320
  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
  4321
  unfolding norm_conv_dist using assms
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4322
  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
  4323
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4324
lemma (in bounded_linear) uniformly_continuous_on:
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4325
  assumes "uniformly_continuous_on s g"
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4326
  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
  4327
  using assms unfolding uniformly_continuous_on_sequentially
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4328
  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
  4329
  by (auto intro: tendsto_zero)
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4330
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4331
lemma uniformly_continuous_on_cmul:
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4332
  fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4333
  assumes "uniformly_continuous_on s f"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4334
  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
  4335
  using bounded_linear_scaleR_right assms
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4336
  by (rule bounded_linear.uniformly_continuous_on)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4337
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4338
lemma dist_minus:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4339
  fixes x y :: "'a::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4340
  shows "dist (- x) (- y) = dist x y"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4341
  unfolding dist_norm minus_diff_minus norm_minus_cancel ..
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4342
44648
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4343
lemma uniformly_continuous_on_minus:
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4344
  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
  4345
  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
  4346
  unfolding uniformly_continuous_on_def dist_minus .
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4347
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4348
lemma uniformly_continuous_on_add:
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4349
  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
  4350
  assumes "uniformly_continuous_on s f"
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4351
  assumes "uniformly_continuous_on s g"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4352
  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
  4353
  using assms unfolding uniformly_continuous_on_sequentially
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4354
  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
  4355
  by (auto intro: tendsto_add_zero)
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4356
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4357
lemma uniformly_continuous_on_diff:
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4358
  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
  4359
  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
  4360
  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
  4361
  unfolding ab_diff_minus using assms
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4362
  by (intro uniformly_continuous_on_add uniformly_continuous_on_minus)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4363
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4364
text{* Continuity of all kinds is preserved under composition. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4365
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4366
lemma continuous_within_topological:
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4367
  "continuous (at x within s) f \<longleftrightarrow>
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4368
    (\<forall>B. open B \<longrightarrow> f x \<in> B \<longrightarrow>
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4369
      (\<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
  4370
unfolding continuous_within
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4371
unfolding tendsto_def Limits.eventually_within eventually_at_topological
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4372
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
  4373
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4374
lemma continuous_within_compose:
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4375
  assumes "continuous (at x within s) f"
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4376
  assumes "continuous (at (f x) within f ` s) g"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4377
  shows "continuous (at x within s) (g o f)"
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4378
using assms unfolding continuous_within_topological by simp metis
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4379
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4380
lemma continuous_at_compose:
45031
9583f2b56f85 add lemmas within_empty and tendsto_bot;
huffman
parents: 44909
diff changeset
  4381
  assumes "continuous (at x) f" and "continuous (at (f x)) g"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4382
  shows "continuous (at x) (g o f)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4383
proof-
45031
9583f2b56f85 add lemmas within_empty and tendsto_bot;
huffman
parents: 44909
diff changeset
  4384
  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
  4385
    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
  4386
  thus ?thesis using assms(1)
9583f2b56f85 add lemmas within_empty and tendsto_bot;
huffman
parents: 44909
diff changeset
  4387
    using continuous_within_compose[of x UNIV f g] by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4388
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4389
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4390
lemma continuous_on_compose:
36440
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  4391
  "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
  4392
  unfolding continuous_on_topological by simp metis
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4393
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4394
lemma uniformly_continuous_on_compose:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4395
  assumes "uniformly_continuous_on s f"  "uniformly_continuous_on (f ` s) g"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4396
  shows "uniformly_continuous_on s (g o f)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4397
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4398
  { fix e::real assume "e>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4399
    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
  4400
    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
  4401
    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
  4402
  thus ?thesis using assms unfolding uniformly_continuous_on_def 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
44647
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4405
lemmas continuous_on_intros = continuous_on_id continuous_on_const
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4406
  continuous_on_compose continuous_on_norm continuous_on_infnorm
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4407
  continuous_on_add continuous_on_minus continuous_on_diff
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4408
  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
  4409
  continuous_on_inner
44648
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4410
  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
  4411
  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
  4412
  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
  4413
  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
  4414
  uniformly_continuous_on_cmul
44647
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4415
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4416
text{* Continuity in terms of open preimages. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4417
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4418
lemma continuous_at_open:
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4419
  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
  4420
unfolding continuous_within_topological [of x UNIV f, unfolded within_UNIV]
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4421
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
  4422
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4423
lemma continuous_on_open:
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4424
  shows "continuous_on s f \<longleftrightarrow>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4425
        (\<forall>t. openin (subtopology euclidean (f ` s)) t
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4426
            --> 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
  4427
proof (safe)
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4428
  fix t :: "'b set"
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4429
  assume 1: "continuous_on s f"
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4430
  assume 2: "openin (subtopology euclidean (f ` s)) t"
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4431
  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
  4432
    unfolding openin_open by auto
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4433
  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
  4434
  have "open U" unfolding U_def by (simp add: open_Union)
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4435
  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
  4436
  proof (intro ballI iffI)
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4437
    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
  4438
      unfolding U_def t by auto
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4439
  next
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4440
    fix x assume "x \<in> s" and "f x \<in> t"
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4441
    hence "x \<in> s" and "f x \<in> B"
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4442
      unfolding t by auto
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4443
    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
  4444
      unfolding t continuous_on_topological by metis
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4445
    then show "x \<in> U"
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4446
      unfolding U_def by auto
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4447
  qed
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4448
  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
  4449
  then show "openin (subtopology euclidean s) {x \<in> s. f x \<in> t}"
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4450
    unfolding openin_open by fast
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4451
next
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4452
  assume "?rhs" show "continuous_on s f"
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4453
  unfolding continuous_on_topological
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4454
  proof (clarify)
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4455
    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
  4456
    have "openin (subtopology euclidean (f ` s)) (f ` s \<inter> B)"
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4457
      unfolding openin_open using `open B` by auto
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4458
    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
  4459
      using `?rhs` by fast
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4460
    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
  4461
      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
  4462
  qed
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4463
qed
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4464
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4465
text {* Similarly in terms of closed sets. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4466
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4467
lemma continuous_on_closed:
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  4468
  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
  4469
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4470
  assume ?lhs
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4471
  { fix t
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4472
    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
  4473
    have **:"f ` s - (f ` s - (f ` s - t)) = f ` s - t" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4474
    assume as:"closedin (subtopology euclidean (f ` s)) t"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4475
    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
  4476
    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
  4477
      unfolding openin_closedin_eq topspace_euclidean_subtopology unfolding * by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4478
  thus ?rhs by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4479
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4480
  assume ?rhs
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4481
  { fix t
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4482
    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
  4483
    assume as:"openin (subtopology euclidean (f ` s)) t"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4484
    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
  4485
      unfolding openin_closedin_eq topspace_euclidean_subtopology *[THEN sym] closedin_subtopology by auto }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4486
  thus ?lhs unfolding continuous_on_open 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
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  4489
text {* Half-global and completely global cases. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4490
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4491
lemma continuous_open_in_preimage:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4492
  assumes "continuous_on s f"  "open t"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4493
  shows "openin (subtopology euclidean s) {x \<in> s. f x \<in> t}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4494
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4495
  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
  4496
  have "openin (subtopology euclidean (f ` s)) (t \<inter> f ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4497
    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
  4498
  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
  4499
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4500
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4501
lemma continuous_closed_in_preimage:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4502
  assumes "continuous_on s f"  "closed t"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4503
  shows "closedin (subtopology euclidean s) {x \<in> s. f x \<in> t}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4504
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4505
  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
  4506
  have "closedin (subtopology euclidean (f ` s)) (t \<inter> f ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4507
    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
  4508
  thus ?thesis
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4509
    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
  4510
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4511
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4512
lemma continuous_open_preimage:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4513
  assumes "continuous_on s f" "open s" "open t"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4514
  shows "open {x \<in> s. f x \<in> t}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4515
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4516
  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
  4517
    using continuous_open_in_preimage[OF assms(1,3)] unfolding openin_open by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4518
  thus ?thesis using open_Int[of s T, OF assms(2)] by auto
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_closed_preimage:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4522
  assumes "continuous_on s f" "closed s" "closed t"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4523
  shows "closed {x \<in> s. f x \<in> t}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4524
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4525
  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
  4526
    using continuous_closed_in_preimage[OF assms(1,3)] unfolding closedin_closed by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4527
  thus ?thesis using closed_Int[of s T, OF assms(2)] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4528
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4529
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4530
lemma continuous_open_preimage_univ:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4531
  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
  4532
  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
  4533
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4534
lemma continuous_closed_preimage_univ:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4535
  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
  4536
  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
  4537
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4538
lemma continuous_open_vimage:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4539
  shows "\<forall>x. continuous (at x) f \<Longrightarrow> open s \<Longrightarrow> open (f -` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4540
  unfolding vimage_def by (rule continuous_open_preimage_univ)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4541
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4542
lemma continuous_closed_vimage:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4543
  shows "\<forall>x. continuous (at x) f \<Longrightarrow> closed s \<Longrightarrow> closed (f -` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4544
  unfolding vimage_def by (rule continuous_closed_preimage_univ)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4545
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4546
lemma interior_image_subset:
35172
579dd5570f96 Added integration to Multivariate-Analysis (upto FTC)
himmelma
parents: 35028
diff changeset
  4547
  assumes "\<forall>x. continuous (at x) f" "inj f"
579dd5570f96 Added integration to Multivariate-Analysis (upto FTC)
himmelma
parents: 35028
diff changeset
  4548
  shows "interior (f ` s) \<subseteq> f ` (interior s)"
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  4549
proof
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  4550
  fix x assume "x \<in> interior (f ` s)"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  4551
  then obtain T where as: "open T" "x \<in> T" "T \<subseteq> f ` s" ..
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  4552
  hence "x \<in> f ` s" by auto
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  4553
  then obtain y where y: "y \<in> s" "x = f y" by auto
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  4554
  have "open (vimage f T)"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  4555
    using assms(1) `open T` by (rule continuous_open_vimage)
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  4556
  moreover have "y \<in> vimage f T"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  4557
    using `x = f y` `x \<in> T` by simp
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  4558
  moreover have "vimage f T \<subseteq> s"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  4559
    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
  4560
  ultimately have "y \<in> interior s" ..
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  4561
  with `x = f y` show "x \<in> f ` interior s" ..
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  4562
qed
35172
579dd5570f96 Added integration to Multivariate-Analysis (upto FTC)
himmelma
parents: 35028
diff changeset
  4563
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  4564
text {* Equality of continuous functions on closure and related results. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4565
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4566
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
  4567
  fixes f :: "_ \<Rightarrow> 'b::t1_space"
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  4568
  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
  4569
  using continuous_closed_in_preimage[of s f "{a}"] by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4570
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4571
lemma continuous_closed_preimage_constant:
36668
941ba2da372e simplify definition of t1_space; generalize lemma closed_sing and related lemmas
huffman
parents: 36667
diff changeset
  4572
  fixes f :: "_ \<Rightarrow> 'b::t1_space"
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  4573
  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
  4574
  using continuous_closed_preimage[of s f "{a}"] by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4575
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4576
lemma continuous_constant_on_closure:
36668
941ba2da372e simplify definition of t1_space; generalize lemma closed_sing and related lemmas
huffman
parents: 36667
diff changeset
  4577
  fixes f :: "_ \<Rightarrow> 'b::t1_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4578
  assumes "continuous_on (closure s) f"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4579
          "\<forall>x \<in> s. f x = a"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4580
  shows "\<forall>x \<in> (closure s). f x = a"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4581
    using continuous_closed_preimage_constant[of "closure s" f a]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4582
    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
  4583
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4584
lemma image_closure_subset:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4585
  assumes "continuous_on (closure s) f"  "closed t"  "(f ` s) \<subseteq> t"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4586
  shows "f ` (closure s) \<subseteq> t"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4587
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4588
  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
  4589
  moreover have "closed {x \<in> closure s. f x \<in> t}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4590
    using continuous_closed_preimage[OF assms(1)] and assms(2) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4591
  ultimately have "closure s = {x \<in> closure s . f x \<in> t}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4592
    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
  4593
  thus ?thesis by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4594
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4595
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4596
lemma continuous_on_closure_norm_le:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4597
  fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4598
  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
  4599
  shows "norm(f x) \<le> b"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4600
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4601
  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
  4602
  show ?thesis
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4603
    using image_closure_subset[OF assms(1) closed_cball[of 0 b] *] assms(3)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4604
    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
  4605
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4606
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  4607
text {* Making a continuous function avoid some value in a neighbourhood. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4608
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4609
lemma continuous_within_avoid:
50898
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4610
  fixes f :: "'a::metric_space \<Rightarrow> 'b::t1_space"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4611
  assumes "continuous (at x within s) f" and "f x \<noteq> a"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4612
  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
  4613
proof-
50898
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4614
  obtain U where "open U" and "f x \<in> U" and "a \<notin> U"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4615
    using t1_space [OF `f x \<noteq> a`] by fast
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4616
  have "(f ---> f x) (at x within s)"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4617
    using assms(1) by (simp add: continuous_within)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4618
  hence "eventually (\<lambda>y. f y \<in> U) (at x within s)"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4619
    using `open U` and `f x \<in> U`
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4620
    unfolding tendsto_def by fast
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4621
  hence "eventually (\<lambda>y. f y \<noteq> a) (at x within s)"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4622
    using `a \<notin> U` by (fast elim: eventually_mono [rotated])
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4623
  thus ?thesis
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4624
    unfolding Limits.eventually_within Limits.eventually_at
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4625
    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
  4626
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4627
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4628
lemma continuous_at_avoid:
50898
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4629
  fixes f :: "'a::metric_space \<Rightarrow> 'b::t1_space"
45031
9583f2b56f85 add lemmas within_empty and tendsto_bot;
huffman
parents: 44909
diff changeset
  4630
  assumes "continuous (at x) f" and "f x \<noteq> a"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4631
  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
  4632
  using assms continuous_within_avoid[of x UNIV f a] by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4633
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4634
lemma continuous_on_avoid:
50898
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4635
  fixes f :: "'a::metric_space \<Rightarrow> 'b::t1_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4636
  assumes "continuous_on s f"  "x \<in> s"  "f x \<noteq> a"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4637
  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
  4638
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
  4639
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4640
lemma continuous_on_open_avoid:
50898
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4641
  fixes f :: "'a::metric_space \<Rightarrow> 'b::t1_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4642
  assumes "continuous_on s f"  "open s"  "x \<in> s"  "f x \<noteq> a"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4643
  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
  4644
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
  4645
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  4646
text {* Proving a function is constant by proving open-ness of level set. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4647
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4648
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
  4649
  fixes f :: "_ \<Rightarrow> 'b::t1_space"
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  4650
  shows "connected s \<Longrightarrow> continuous_on s f \<Longrightarrow>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4651
        openin (subtopology euclidean s) {x \<in> s. f x = a}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4652
        ==> (\<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
  4653
unfolding connected_clopen using continuous_closed_in_preimage_constant by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4654
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4655
lemma continuous_levelset_open_in:
36668
941ba2da372e simplify definition of t1_space; generalize lemma closed_sing and related lemmas
huffman
parents: 36667
diff changeset
  4656
  fixes f :: "_ \<Rightarrow> 'b::t1_space"
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  4657
  shows "connected s \<Longrightarrow> continuous_on s f \<Longrightarrow>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4658
        openin (subtopology euclidean s) {x \<in> s. f x = a} \<Longrightarrow>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4659
        (\<exists>x \<in> s. f x = a)  ==> (\<forall>x \<in> s. f x = a)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4660
using continuous_levelset_open_in_cases[of s f ]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4661
by meson
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4662
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4663
lemma continuous_levelset_open:
36668
941ba2da372e simplify definition of t1_space; generalize lemma closed_sing and related lemmas
huffman
parents: 36667
diff changeset
  4664
  fixes f :: "_ \<Rightarrow> 'b::t1_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4665
  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
  4666
  shows "\<forall>x \<in> s. f x = a"
36362
06475a1547cb fix lots of looping simp calls and other warnings
huffman
parents: 36360
diff changeset
  4667
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
  4668
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  4669
text {* Some arithmetical combinations (more to prove). *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4670
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4671
lemma open_scaling[intro]:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4672
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4673
  assumes "c \<noteq> 0"  "open s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4674
  shows "open((\<lambda>x. c *\<^sub>R x) ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4675
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4676
  { fix x assume "x \<in> s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4677
    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
  4678
    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
  4679
    moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4680
    { fix y assume "dist y (c *\<^sub>R x) < e * \<bar>c\<bar>"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4681
      hence "norm ((1 / c) *\<^sub>R y - x) < e" unfolding dist_norm
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4682
        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
  4683
          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
  4684
      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
  4685
    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
  4686
  thus ?thesis unfolding open_dist by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4687
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4688
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4689
lemma minus_image_eq_vimage:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4690
  fixes A :: "'a::ab_group_add set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4691
  shows "(\<lambda>x. - x) ` A = (\<lambda>x. - x) -` A"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4692
  by (auto intro!: image_eqI [where f="\<lambda>x. - x"])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4693
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4694
lemma open_negations:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4695
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4696
  shows "open s ==> open ((\<lambda> x. -x) ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4697
  unfolding scaleR_minus1_left [symmetric]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4698
  by (rule open_scaling, auto)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4699
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4700
lemma open_translation:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4701
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4702
  assumes "open s"  shows "open((\<lambda>x. a + x) ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4703
proof-
44647
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4704
  { fix x have "continuous (at x) (\<lambda>x. x - a)"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4705
      by (intro continuous_diff continuous_at_id continuous_const) }
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4706
  moreover have "{x. x - a \<in> s} = op + a ` s" by force
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4707
  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
  4708
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4709
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4710
lemma open_affinity:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4711
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4712
  assumes "open s"  "c \<noteq> 0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4713
  shows "open ((\<lambda>x. a + c *\<^sub>R x) ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4714
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4715
  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
  4716
  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
  4717
  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
  4718
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4719
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4720
lemma interior_translation:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4721
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4722
  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
  4723
proof (rule set_eqI, rule)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4724
  fix x assume "x \<in> interior (op + a ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4725
  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
  4726
  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
  4727
  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
  4728
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4729
  fix x assume "x \<in> op + a ` interior s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4730
  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
  4731
  { fix z have *:"a + y - z = y + a - z" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4732
    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
  4733
    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
  4734
    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
  4735
  hence "ball x e \<subseteq> op + a ` s" unfolding subset_eq by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4736
  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
  4737
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4738
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4739
text {* Topological properties of linear functions. *}
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4740
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4741
lemma linear_lim_0:
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4742
  assumes "bounded_linear f" shows "(f ---> 0) (at (0))"
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4743
proof-
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4744
  interpret f: bounded_linear f by fact
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4745
  have "(f ---> f 0) (at 0)"
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4746
    using tendsto_ident_at by (rule f.tendsto)
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4747
  thus ?thesis unfolding f.zero .
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4748
qed
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4749
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4750
lemma linear_continuous_at:
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4751
  assumes "bounded_linear f"  shows "continuous (at a) f"
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4752
  unfolding continuous_at using assms
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4753
  apply (rule bounded_linear.tendsto)
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4754
  apply (rule tendsto_ident_at)
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4755
  done
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4756
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4757
lemma linear_continuous_within:
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4758
  shows "bounded_linear f ==> continuous (at x within s) f"
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4759
  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
  4760
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4761
lemma linear_continuous_on:
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4762
  shows "bounded_linear f ==> continuous_on s f"
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4763
  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
  4764
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  4765
text {* Also bilinear functions, in composition form. *}
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4766
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4767
lemma bilinear_continuous_at_compose:
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4768
  shows "continuous (at x) f \<Longrightarrow> continuous (at x) g \<Longrightarrow> bounded_bilinear h
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4769
        ==> continuous (at x) (\<lambda>x. h (f x) (g x))"
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4770
  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
  4771
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4772
lemma bilinear_continuous_within_compose:
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4773
  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
  4774
        ==> continuous (at x within s) (\<lambda>x. h (f x) (g x))"
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4775
  unfolding continuous_within using Lim_bilinear[of f "f x"] by auto
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4776
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4777
lemma bilinear_continuous_on_compose:
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4778
  shows "continuous_on s f \<Longrightarrow> continuous_on s g \<Longrightarrow> bounded_bilinear h
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4779
             ==> continuous_on s (\<lambda>x. h (f x) (g x))"
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4780
  unfolding continuous_on_def
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4781
  by (fast elim: bounded_bilinear.tendsto)
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4782
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  4783
text {* Preservation of compactness and connectedness under continuous function. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4784
50898
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4785
lemma compact_eq_openin_cover:
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4786
  "compact S \<longleftrightarrow>
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4787
    (\<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
  4788
      (\<exists>D\<subseteq>C. finite D \<and> S \<subseteq> \<Union>D))"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4789
proof safe
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4790
  fix C
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4791
  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
  4792
  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
  4793
    unfolding openin_open by force+
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4794
  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
  4795
    by (rule compactE)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4796
  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
  4797
    by auto
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4798
  thus "\<exists>D\<subseteq>C. finite D \<and> S \<subseteq> \<Union>D" ..
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4799
next
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4800
  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
  4801
        (\<exists>D\<subseteq>C. finite D \<and> S \<subseteq> \<Union>D)"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4802
  show "compact S"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4803
  proof (rule compactI)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4804
    fix C
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4805
    let ?C = "image (\<lambda>T. S \<inter> T) C"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4806
    assume "\<forall>t\<in>C. open t" and "S \<subseteq> \<Union>C"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4807
    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
  4808
      unfolding openin_open by auto
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4809
    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
  4810
      by metis
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4811
    let ?D = "inv_into C (\<lambda>T. S \<inter> T) ` D"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4812
    have "?D \<subseteq> C \<and> finite ?D \<and> S \<subseteq> \<Union>?D"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4813
    proof (intro conjI)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4814
      from `D \<subseteq> ?C` show "?D \<subseteq> C"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4815
        by (fast intro: inv_into_into)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4816
      from `finite D` show "finite ?D"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4817
        by (rule finite_imageI)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4818
      from `S \<subseteq> \<Union>D` show "S \<subseteq> \<Union>?D"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4819
        apply (rule subset_trans)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4820
        apply clarsimp
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4821
        apply (frule subsetD [OF `D \<subseteq> ?C`, THEN f_inv_into_f])
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4822
        apply (erule rev_bexI, fast)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4823
        done
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4824
    qed
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4825
    thus "\<exists>D\<subseteq>C. finite D \<and> S \<subseteq> \<Union>D" ..
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4826
  qed
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4827
qed
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4828
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4829
lemma compact_continuous_image:
50898
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4830
  assumes "continuous_on s f" and "compact s"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4831
  shows "compact (f ` s)"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4832
using assms (* FIXME: long unstructured proof *)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4833
unfolding continuous_on_open
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4834
unfolding compact_eq_openin_cover
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4835
apply clarify
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4836
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
  4837
apply (drule mp)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4838
apply (rule conjI)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4839
apply simp
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4840
apply clarsimp
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4841
apply (drule subsetD)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4842
apply (erule imageI)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4843
apply fast
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4844
apply (erule thin_rl)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4845
apply clarify
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4846
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
  4847
apply (intro conjI)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4848
apply clarify
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4849
apply (rule inv_into_into)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4850
apply (erule (1) subsetD)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4851
apply (erule finite_imageI)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4852
apply (clarsimp, rename_tac x)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4853
apply (drule (1) subsetD, clarify)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4854
apply (drule (1) subsetD, clarify)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4855
apply (rule rev_bexI)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4856
apply assumption
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4857
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
  4858
apply (drule f_inv_into_f)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4859
apply fast
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4860
apply (erule imageI)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4861
done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4862
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4863
lemma connected_continuous_image:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4864
  assumes "continuous_on s f"  "connected s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4865
  shows "connected(f ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4866
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4867
  { 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
  4868
    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
  4869
      using assms(1)[unfolded continuous_on_open, THEN spec[where x=T]]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4870
      using assms(1)[unfolded continuous_on_closed, THEN spec[where x=T]]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4871
      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
  4872
    hence False using as(1,2)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4873
      using as(4)[unfolded closedin_def topspace_euclidean_subtopology] by auto }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4874
  thus ?thesis unfolding connected_clopen by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4875
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4876
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  4877
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
  4878
  
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4879
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
  4880
  assumes f: "continuous_on s f" and s: "compact s"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4881
  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
  4882
  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
  4883
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
  4884
  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
  4885
  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
  4886
  let ?b = "(\<lambda>(y, d). ball y (d/2))"
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  4887
  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
  4888
  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
  4889
    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
  4890
    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
  4891
    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
  4892
      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
  4893
    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
  4894
      using `0 < e` `y \<in> s` by (auto elim!: openE)
50944
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  4895
    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
  4896
      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
  4897
  qed auto
50944
03b11adf1f33 simplified prove of compact_imp_bounded
hoelzl
parents: 50943
diff changeset
  4898
  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
  4899
    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
  4900
  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
  4901
    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
  4902
  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
  4903
  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
  4904
    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
  4905
    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
  4906
      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
  4907
    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
  4908
    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
  4909
      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
  4910
    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
  4911
      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
  4912
  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
  4913
qed auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4914
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4915
text{* Continuity of inverse function on compact domain. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4916
44647
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4917
lemma continuous_on_inv:
50898
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4918
  fixes f :: "'a::topological_space \<Rightarrow> 'b::t2_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4919
  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
  4920
  shows "continuous_on (f ` s) g"
50898
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4921
unfolding continuous_on_topological
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4922
proof (clarsimp simp add: assms(3))
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4923
  fix x :: 'a and B :: "'a set"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4924
  assume "x \<in> s" and "open B" and "x \<in> B"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4925
  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
  4926
    using assms(3) by (auto, metis)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4927
  have "continuous_on (s - B) f"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4928
    using `continuous_on s f` Diff_subset
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4929
    by (rule continuous_on_subset)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4930
  moreover have "compact (s - B)"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4931
    using `open B` and `compact s`
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4932
    unfolding Diff_eq by (intro compact_inter_closed closed_Compl)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4933
  ultimately have "compact (f ` (s - B))"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4934
    by (rule compact_continuous_image)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4935
  hence "closed (f ` (s - B))"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4936
    by (rule compact_imp_closed)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4937
  hence "open (- f ` (s - B))"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4938
    by (rule open_Compl)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4939
  moreover have "f x \<in> - f ` (s - B)"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4940
    using `x \<in> s` and `x \<in> B` by (simp add: 1)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4941
  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
  4942
    by (simp add: 1)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4943
  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
  4944
    by fast
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4945
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4946
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4947
text {* A uniformly convergent limit of continuous functions is continuous. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4948
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4949
lemma continuous_uniform_limit:
44212
4d10e7f342b1 generalize lemma continuous_uniform_limit to class metric_space
huffman
parents: 44211
diff changeset
  4950
  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
  4951
  assumes "\<not> trivial_limit F"
4d10e7f342b1 generalize lemma continuous_uniform_limit to class metric_space
huffman
parents: 44211
diff changeset
  4952
  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
  4953
  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
  4954
  shows "continuous_on s g"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4955
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4956
  { 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
  4957
    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
  4958
      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
  4959
    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
  4960
    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
  4961
      using assms(1) by blast
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4962
    have "e / 3 > 0" using `e>0` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4963
    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
  4964
      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
  4965
    { 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
  4966
      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
  4967
        by (rule d [rule_format])
4d10e7f342b1 generalize lemma continuous_uniform_limit to class metric_space
huffman
parents: 44211
diff changeset
  4968
      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
  4969
        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
  4970
        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
  4971
        by auto
4d10e7f342b1 generalize lemma continuous_uniform_limit to class metric_space
huffman
parents: 44211
diff changeset
  4972
      hence "dist (g y) (g x) < e"
4d10e7f342b1 generalize lemma continuous_uniform_limit to class metric_space
huffman
parents: 44211
diff changeset
  4973
        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
  4974
        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
  4975
        by auto }
4d10e7f342b1 generalize lemma continuous_uniform_limit to class metric_space
huffman
parents: 44211
diff changeset
  4976
    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
  4977
      using `d>0` by auto }
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  4978
  thus ?thesis unfolding continuous_on_iff by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4979
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4980
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  4981
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  4982
subsection {* Topological stuff lifted from and dropped to R *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4983
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4984
lemma open_real:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4985
  fixes s :: "real set" shows
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4986
 "open s \<longleftrightarrow>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4987
        (\<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
  4988
  unfolding open_dist dist_norm by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4989
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4990
lemma islimpt_approachable_real:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4991
  fixes s :: "real set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4992
  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
  4993
  unfolding islimpt_approachable dist_norm by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4994
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4995
lemma closed_real:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4996
  fixes s :: "real set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4997
  shows "closed s \<longleftrightarrow>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4998
        (\<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
  4999
            --> x \<in> s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5000
  unfolding closed_limpt islimpt_approachable dist_norm by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5001
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5002
lemma continuous_at_real_range:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5003
  fixes f :: "'a::real_normed_vector \<Rightarrow> real"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5004
  shows "continuous (at x) f \<longleftrightarrow> (\<forall>e>0. \<exists>d>0.
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5005
        \<forall>x'. norm(x' - x) < d --> abs(f x' - f x) < e)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5006
  unfolding continuous_at unfolding Lim_at
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5007
  unfolding dist_nz[THEN sym] unfolding dist_norm apply auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5008
  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
  5009
  apply(erule_tac x=e in allE) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5010
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5011
lemma continuous_on_real_range:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5012
  fixes f :: "'a::real_normed_vector \<Rightarrow> real"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5013
  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
  5014
  unfolding continuous_on_iff dist_norm by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5015
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  5016
text {* Hence some handy theorems on distance, diameter etc. of/from a set. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5017
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5018
lemma compact_attains_sup:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5019
  fixes s :: "real set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5020
  assumes "compact s"  "s \<noteq> {}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5021
  shows "\<exists>x \<in> s. \<forall>y \<in> s. y \<le> x"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5022
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5023
  from assms(1) have a:"bounded s" "closed s" unfolding compact_eq_bounded_closed by auto
33270
paulson
parents: 33175
diff changeset
  5024
  { 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
  5025
    have "isLub UNIV s (Sup s)" using Sup[OF assms(2)] unfolding setle_def using as(1) by auto
paulson
parents: 33175
diff changeset
  5026
    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
  5027
    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
  5028
  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
  5029
    apply(rule_tac x="Sup s" in bexI) by auto
paulson
parents: 33175
diff changeset
  5030
qed
paulson
parents: 33175
diff changeset
  5031
paulson
parents: 33175
diff changeset
  5032
lemma Inf:
paulson
parents: 33175
diff changeset
  5033
  fixes S :: "real set"
paulson
parents: 33175
diff changeset
  5034
  shows "S \<noteq> {} ==> (\<exists>b. b <=* S) ==> isGlb UNIV S (Inf S)"
paulson
parents: 33175
diff changeset
  5035
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
  5036
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5037
lemma compact_attains_inf:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5038
  fixes s :: "real set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5039
  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
  5040
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5041
  from assms(1) have a:"bounded s" "closed s" unfolding compact_eq_bounded_closed by auto
33270
paulson
parents: 33175
diff changeset
  5042
  { fix e::real assume as: "\<forall>x\<in>s. x \<ge> Inf s"  "Inf s \<notin> s"  "0 < e"
paulson
parents: 33175
diff changeset
  5043
      "\<forall>x'\<in>s. x' = Inf s \<or> \<not> abs (x' - Inf s) < e"
paulson
parents: 33175
diff changeset
  5044
    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
  5045
    moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5046
    { fix x assume "x \<in> s"
33270
paulson
parents: 33175
diff changeset
  5047
      hence *:"abs (x - Inf s) = x - Inf s" using as(1)[THEN bspec[where x=x]] by auto
paulson
parents: 33175
diff changeset
  5048
      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
  5049
    hence "isLb UNIV s (Inf s + e)" unfolding isLb_def and setge_def by auto
paulson
parents: 33175
diff changeset
  5050
    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
  5051
  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
  5052
    apply(rule_tac x="Inf s" in bexI) by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5053
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5054
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5055
lemma continuous_attains_sup:
50948
8c742f9de9f5 generalize topology lemmas; simplify proofs
huffman
parents: 50944
diff changeset
  5056
  fixes f :: "'a::topological_space \<Rightarrow> real"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5057
  shows "compact s \<Longrightarrow> s \<noteq> {} \<Longrightarrow> continuous_on s f
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5058
        ==> (\<exists>x \<in> s. \<forall>y \<in> s.  f y \<le> f x)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5059
  using compact_attains_sup[of "f ` s"]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5060
  using compact_continuous_image[of s f] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5061
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5062
lemma continuous_attains_inf:
50948
8c742f9de9f5 generalize topology lemmas; simplify proofs
huffman
parents: 50944
diff changeset
  5063
  fixes f :: "'a::topological_space \<Rightarrow> real"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5064
  shows "compact s \<Longrightarrow> s \<noteq> {} \<Longrightarrow> continuous_on s f
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5065
        \<Longrightarrow> (\<exists>x \<in> s. \<forall>y \<in> s. f x \<le> f y)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5066
  using compact_attains_inf[of "f ` s"]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5067
  using compact_continuous_image[of s f] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5068
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5069
lemma distance_attains_sup:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5070
  assumes "compact s" "s \<noteq> {}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5071
  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
  5072
proof (rule continuous_attains_sup [OF assms])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5073
  { fix x assume "x\<in>s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5074
    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
  5075
      by (intro tendsto_dist tendsto_const Lim_at_within tendsto_ident_at)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5076
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5077
  thus "continuous_on s (dist a)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5078
    unfolding continuous_on ..
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5079
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5080
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  5081
text {* For \emph{minimal} distance, we only need closure, not compactness. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5082
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5083
lemma distance_attains_inf:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5084
  fixes a :: "'a::heine_borel"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5085
  assumes "closed s"  "s \<noteq> {}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5086
  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
  5087
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5088
  from assms(2) obtain b where "b\<in>s" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5089
  let ?B = "cball a (dist b a) \<inter> s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5090
  have "b \<in> ?B" using `b\<in>s` by (simp add: dist_commute)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5091
  hence "?B \<noteq> {}" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5092
  moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5093
  { fix x assume "x\<in>?B"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5094
    fix e::real assume "e>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5095
    { fix x' assume "x'\<in>?B" and as:"dist x' x < e"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5096
      from as have "\<bar>dist a x' - dist a x\<bar> < e"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5097
        unfolding abs_less_iff minus_diff_eq
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5098
        using dist_triangle2 [of a x' x]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5099
        using dist_triangle [of a x x']
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5100
        by arith
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5101
    }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5102
    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
  5103
      using `e>0` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5104
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5105
  hence "continuous_on (cball a (dist b a) \<inter> s) (dist a)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5106
    unfolding continuous_on Lim_within dist_norm real_norm_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5107
    by fast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5108
  moreover have "compact ?B"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5109
    using compact_cball[of a "dist b a"]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5110
    unfolding compact_eq_bounded_closed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5111
    using bounded_Int and closed_Int and assms(1) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5112
  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
  5113
    using continuous_attains_inf[of ?B "dist a"] by fastforce
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 44668
diff changeset
  5114
  thus ?thesis by fastforce
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5115
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5116
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  5117
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  5118
subsection {* Pasted sets *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5119
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5120
lemma bounded_Times:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5121
  assumes "bounded s" "bounded t" shows "bounded (s \<times> t)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5122
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5123
  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
  5124
    using assms [unfolded bounded_def] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5125
  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
  5126
    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
  5127
  thus ?thesis unfolding bounded_any_center [where a="(x, y)"] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5128
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5129
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5130
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
  5131
by (induct x) simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5132
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  5133
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
  5134
unfolding seq_compact_def
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5135
apply clarify
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5136
apply (drule_tac x="fst \<circ> f" in spec)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5137
apply (drule mp, simp add: mem_Times_iff)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5138
apply (clarify, rename_tac l1 r1)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5139
apply (drule_tac x="snd \<circ> f \<circ> r1" in spec)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5140
apply (drule mp, simp add: mem_Times_iff)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5141
apply (clarify, rename_tac l2 r2)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5142
apply (rule_tac x="(l1, l2)" in rev_bexI, simp)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5143
apply (rule_tac x="r1 \<circ> r2" in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5144
apply (rule conjI, simp add: subseq_def)
48125
602dc0215954 tuned proofs -- prefer direct "rotated" instead of old-style COMP;
wenzelm
parents: 48048
diff changeset
  5145
apply (drule_tac r=r2 in lim_subseq [rotated], assumption)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5146
apply (drule (1) tendsto_Pair) back
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5147
apply (simp add: o_def)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5148
done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5149
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  5150
text {* Generalize to @{class topological_space} *}
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  5151
lemma compact_Times: 
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  5152
  fixes s :: "'a::metric_space set" and t :: "'b::metric_space set"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  5153
  shows "compact s \<Longrightarrow> compact t \<Longrightarrow> compact (s \<times> t)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  5154
  unfolding compact_eq_seq_compact_metric by (rule seq_compact_Times)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  5155
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  5156
text{* Hence some useful properties follow quite easily. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5157
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5158
lemma compact_scaling:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5159
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5160
  assumes "compact s"  shows "compact ((\<lambda>x. c *\<^sub>R x) ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5161
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5162
  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
  5163
  have *:"bounded_linear ?f" by (rule bounded_linear_scaleR_right)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5164
  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
  5165
    using linear_continuous_at[OF *] assms by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5166
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5167
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5168
lemma compact_negations:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5169
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5170
  assumes "compact s"  shows "compact ((\<lambda>x. -x) ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5171
  using compact_scaling [OF assms, of "- 1"] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5172
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5173
lemma compact_sums:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5174
  fixes s t :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5175
  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
  5176
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5177
  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
  5178
    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
  5179
  have "continuous_on (s \<times> t) (\<lambda>z. fst z + snd z)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5180
    unfolding continuous_on by (rule ballI) (intro tendsto_intros)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5181
  thus ?thesis unfolding * using compact_continuous_image compact_Times [OF assms] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5182
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5183
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5184
lemma compact_differences:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5185
  fixes s t :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5186
  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
  5187
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5188
  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
  5189
    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
  5190
  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
  5191
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5192
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5193
lemma compact_translation:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5194
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5195
  assumes "compact s"  shows "compact ((\<lambda>x. a + x) ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5196
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5197
  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
  5198
  thus ?thesis using compact_sums[OF assms compact_sing[of a]] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5199
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5200
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5201
lemma compact_affinity:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5202
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5203
  assumes "compact s"  shows "compact ((\<lambda>x. a + c *\<^sub>R x) ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5204
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5205
  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
  5206
  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
  5207
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5208
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  5209
text {* Hence we get the following. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5210
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5211
lemma compact_sup_maxdistance:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5212
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5213
  assumes "compact s"  "s \<noteq> {}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5214
  shows "\<exists>x\<in>s. \<exists>y\<in>s. \<forall>u\<in>s. \<forall>v\<in>s. norm(u - v) \<le> norm(x - y)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5215
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5216
  have "{x - y | x y . x\<in>s \<and> y\<in>s} \<noteq> {}" using `s \<noteq> {}` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5217
  then obtain x where x:"x\<in>{x - y |x y. x \<in> s \<and> y \<in> s}"  "\<forall>y\<in>{x - y |x y. x \<in> s \<and> y \<in> s}. norm y \<le> norm x"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5218
    using compact_differences[OF assms(1) assms(1)]
36362
06475a1547cb fix lots of looping simp calls and other warnings
huffman
parents: 36360
diff changeset
  5219
    using distance_attains_sup[where 'a="'a", unfolded dist_norm, of "{x - y | x y . x\<in>s \<and> y\<in>s}" 0] by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5220
  from x(1) obtain a b where "a\<in>s" "b\<in>s" "x = a - b" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5221
  thus ?thesis using x(2)[unfolded `x = a - b`] by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5222
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5223
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  5224
text {* We can state this in terms of diameter of a set. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5225
33270
paulson
parents: 33175
diff changeset
  5226
definition "diameter s = (if s = {} then 0::real else Sup {norm(x - y) | x y. x \<in> s \<and> y \<in> s})"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5227
  (* TODO: generalize to class metric_space *)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5228
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5229
lemma diameter_bounded:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5230
  assumes "bounded s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5231
  shows "\<forall>x\<in>s. \<forall>y\<in>s. norm(x - y) \<le> diameter s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5232
        "\<forall>d>0. d < diameter s --> (\<exists>x\<in>s. \<exists>y\<in>s. norm(x - y) > d)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5233
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5234
  let ?D = "{norm (x - y) |x y. x \<in> s \<and> y \<in> s}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5235
  obtain a where a:"\<forall>x\<in>s. norm x \<le> a" using assms[unfolded bounded_iff] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5236
  { fix x y assume "x \<in> s" "y \<in> s"
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36336
diff changeset
  5237
    hence "norm (x - y) \<le> 2 * a" using norm_triangle_ineq[of x "-y", unfolded norm_minus_cancel] a[THEN bspec[where x=x]] a[THEN bspec[where x=y]] by (auto simp add: field_simps)  }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5238
  note * = this
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5239
  { fix x y assume "x\<in>s" "y\<in>s"  hence "s \<noteq> {}" by auto
36362
06475a1547cb fix lots of looping simp calls and other warnings
huffman
parents: 36360
diff changeset
  5240
    have "norm(x - y) \<le> diameter s" unfolding diameter_def using `s\<noteq>{}` *[OF `x\<in>s` `y\<in>s`] `x\<in>s` `y\<in>s`
44584
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  5241
      by simp (blast del: Sup_upper intro!: * Sup_upper) }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5242
  moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5243
  { fix d::real assume "d>0" "d < diameter s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5244
    hence "s\<noteq>{}" unfolding diameter_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5245
    have "\<exists>d' \<in> ?D. d' > d"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5246
    proof(rule ccontr)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5247
      assume "\<not> (\<exists>d'\<in>{norm (x - y) |x y. x \<in> s \<and> y \<in> s}. d < d')"
33324
51eb2ffa2189 Tidied up some very ugly proofs
paulson
parents: 33270
diff changeset
  5248
      hence "\<forall>d'\<in>?D. d' \<le> d" by auto (metis not_leE) 
51eb2ffa2189 Tidied up some very ugly proofs
paulson
parents: 33270
diff changeset
  5249
      thus False using `d < diameter s` `s\<noteq>{}` 
51eb2ffa2189 Tidied up some very ugly proofs
paulson
parents: 33270
diff changeset
  5250
        apply (auto simp add: diameter_def) 
51eb2ffa2189 Tidied up some very ugly proofs
paulson
parents: 33270
diff changeset
  5251
        apply (drule Sup_real_iff [THEN [2] rev_iffD2])
51eb2ffa2189 Tidied up some very ugly proofs
paulson
parents: 33270
diff changeset
  5252
        apply (auto, force) 
51eb2ffa2189 Tidied up some very ugly proofs
paulson
parents: 33270
diff changeset
  5253
        done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5254
    qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5255
    hence "\<exists>x\<in>s. \<exists>y\<in>s. norm(x - y) > d" by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5256
  ultimately show "\<forall>x\<in>s. \<forall>y\<in>s. norm(x - y) \<le> diameter s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5257
        "\<forall>d>0. d < diameter s --> (\<exists>x\<in>s. \<exists>y\<in>s. norm(x - y) > d)" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5258
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5259
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5260
lemma diameter_bounded_bound:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5261
 "bounded s \<Longrightarrow> x \<in> s \<Longrightarrow> y \<in> s ==> norm(x - y) \<le> diameter s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5262
  using diameter_bounded by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5263
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5264
lemma diameter_compact_attained:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5265
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5266
  assumes "compact s"  "s \<noteq> {}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5267
  shows "\<exists>x\<in>s. \<exists>y\<in>s. (norm(x - y) = diameter s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5268
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5269
  have b:"bounded s" using assms(1) by (rule compact_imp_bounded)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5270
  then obtain x y where xys:"x\<in>s" "y\<in>s" and xy:"\<forall>u\<in>s. \<forall>v\<in>s. norm (u - v) \<le> norm (x - y)" using compact_sup_maxdistance[OF assms] by auto
36362
06475a1547cb fix lots of looping simp calls and other warnings
huffman
parents: 36360
diff changeset
  5271
  hence "diameter s \<le> norm (x - y)"
06475a1547cb fix lots of looping simp calls and other warnings
huffman
parents: 36360
diff changeset
  5272
    unfolding diameter_def by clarsimp (rule Sup_least, fast+)
33324
51eb2ffa2189 Tidied up some very ugly proofs
paulson
parents: 33270
diff changeset
  5273
  thus ?thesis
36362
06475a1547cb fix lots of looping simp calls and other warnings
huffman
parents: 36360
diff changeset
  5274
    by (metis b diameter_bounded_bound order_antisym xys)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5275
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5276
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  5277
text {* Related results with closure as the conclusion. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5278
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5279
lemma closed_scaling:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5280
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5281
  assumes "closed s" shows "closed ((\<lambda>x. c *\<^sub>R x) ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5282
proof(cases "s={}")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5283
  case True thus ?thesis by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5284
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5285
  case False
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5286
  show ?thesis
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5287
  proof(cases "c=0")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5288
    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
  5289
    case True thus ?thesis apply auto unfolding * by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5290
  next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5291
    case False
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5292
    { 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
  5293
      { fix n::nat have "scaleR (1 / c) (x n) \<in> s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5294
          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
  5295
          using `c\<noteq>0` by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5296
      }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5297
      moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5298
      { fix e::real assume "e>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5299
        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
  5300
        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
  5301
          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
  5302
        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
  5303
          unfolding dist_norm unfolding scaleR_right_diff_distrib[THEN sym]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5304
          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
  5305
      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
  5306
      ultimately have "l \<in> scaleR c ` s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5307
        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
  5308
        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
  5309
    thus ?thesis unfolding closed_sequential_limits by fast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5310
  qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5311
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5312
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5313
lemma closed_negations:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5314
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5315
  assumes "closed s"  shows "closed ((\<lambda>x. -x) ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5316
  using closed_scaling[OF assms, of "- 1"] by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5317
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5318
lemma compact_closed_sums:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5319
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5320
  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
  5321
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5322
  let ?S = "{x + y |x y. x \<in> s \<and> y \<in> t}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5323
  { fix x l assume as:"\<forall>n. x n \<in> ?S"  "(x ---> l) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5324
    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
  5325
      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
  5326
    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
  5327
      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
  5328
    have "((\<lambda>n. snd (f (r n))) ---> l - l') sequentially"
44125
230a8665c919 mark some redundant theorems as legacy
huffman
parents: 44122
diff changeset
  5329
      using tendsto_diff[OF lim_subseq[OF r as(2)] lr] and f(1) unfolding o_def by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5330
    hence "l - l' \<in> t"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5331
      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
  5332
      using f(3) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5333
    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
  5334
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5335
  thus ?thesis unfolding closed_sequential_limits by fast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5336
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5337
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5338
lemma closed_compact_sums:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5339
  fixes s t :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5340
  assumes "closed s"  "compact t"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5341
  shows "closed {x + y | x y. x \<in> s \<and> y \<in> t}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5342
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5343
  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
  5344
    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
  5345
  thus ?thesis using compact_closed_sums[OF assms(2,1)] by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5346
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5347
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5348
lemma compact_closed_differences:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5349
  fixes s t :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5350
  assumes "compact s"  "closed t"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5351
  shows "closed {x - y | x y. x \<in> s \<and> y \<in> t}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5352
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5353
  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
  5354
    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
  5355
  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
  5356
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5357
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5358
lemma closed_compact_differences:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5359
  fixes s t :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5360
  assumes "closed s" "compact t"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5361
  shows "closed {x - y | x y. x \<in> s \<and> y \<in> t}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5362
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5363
  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
  5364
    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
  5365
 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
  5366
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5367
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5368
lemma closed_translation:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5369
  fixes a :: "'a::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5370
  assumes "closed s"  shows "closed ((\<lambda>x. a + x) ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5371
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5372
  have "{a + y |y. y \<in> s} = (op + a ` s)" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5373
  thus ?thesis using compact_closed_sums[OF compact_sing[of a] assms] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5374
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5375
34105
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
  5376
lemma translation_Compl:
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
  5377
  fixes a :: "'a::ab_group_add"
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
  5378
  shows "(\<lambda>x. a + x) ` (- t) = - ((\<lambda>x. a + x) ` t)"
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
  5379
  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
  5380
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5381
lemma translation_UNIV:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5382
  fixes a :: "'a::ab_group_add" shows "range (\<lambda>x. a + x) = UNIV"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5383
  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
  5384
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5385
lemma translation_diff:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5386
  fixes a :: "'a::ab_group_add"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5387
  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
  5388
  by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5389
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5390
lemma closure_translation:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5391
  fixes a :: "'a::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5392
  shows "closure ((\<lambda>x. a + x) ` s) = (\<lambda>x. a + x) ` (closure s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5393
proof-
34105
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
  5394
  have *:"op + a ` (- s) = - op + a ` s"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5395
    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
  5396
  show ?thesis unfolding closure_interior translation_Compl
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
  5397
    using interior_translation[of a "- s"] unfolding * by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5398
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5399
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5400
lemma frontier_translation:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5401
  fixes a :: "'a::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5402
  shows "frontier((\<lambda>x. a + x) ` s) = (\<lambda>x. a + x) ` (frontier s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5403
  unfolding frontier_def translation_diff interior_translation closure_translation by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5404
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  5405
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  5406
subsection {* Separation between points and sets *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5407
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5408
lemma separate_point_closed:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5409
  fixes s :: "'a::heine_borel set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5410
  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
  5411
proof(cases "s = {}")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5412
  case True
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5413
  thus ?thesis by(auto intro!: exI[where x=1])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5414
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5415
  case False
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5416
  assume "closed s" "a \<notin> s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5417
  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
  5418
  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
  5419
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5420
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5421
lemma separate_compact_closed:
50949
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5422
  fixes s t :: "'a::heine_borel set"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5423
  assumes "compact s" and "closed t" and "s \<inter> t = {}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5424
  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
  5425
proof - (* FIXME: long proof *)
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5426
  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
  5427
  note `compact s`
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5428
  moreover have "\<forall>t\<in>?T. open t" by auto
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5429
  moreover have "s \<subseteq> \<Union>?T"
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5430
    apply auto
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5431
    apply (rule rev_bexI, assumption)
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5432
    apply (subgoal_tac "x \<notin> t")
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5433
    apply (drule separate_point_closed [OF `closed t`])
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5434
    apply clarify
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5435
    apply (rule_tac x="ball x (d / 2)" in exI)
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5436
    apply simp
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5437
    apply fast
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5438
    apply (cut_tac assms(3))
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5439
    apply auto
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5440
    done
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5441
  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
  5442
    by (rule compactE)
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5443
  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
  5444
    apply (induct set: finite)
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5445
    apply simp
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5446
    apply (rule exI)
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5447
    apply (rule zero_less_one)
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5448
    apply clarsimp
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5449
    apply (rename_tac y e)
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5450
    apply (rule_tac x="min d (e / 2)" in exI)
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5451
    apply simp
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5452
    apply (subst ball_Un)
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5453
    apply (rule conjI)
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5454
    apply (intro ballI, rename_tac z)
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5455
    apply (rule min_max.le_infI2)
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5456
    apply (simp only: mem_ball)
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5457
    apply (drule (1) bspec)
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5458
    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
  5459
    apply simp
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5460
    apply (intro ballI)
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5461
    apply (rule min_max.le_infI1)
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5462
    apply simp
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5463
    done
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5464
  with `s \<subseteq> \<Union>U` show ?thesis
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5465
    by fast
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5466
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5467
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5468
lemma separate_closed_compact:
50949
a5689bb4ed7e generalize more topology lemmas
huffman
parents: 50948
diff changeset
  5469
  fixes s t :: "'a::heine_borel set"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5470
  assumes "closed s" and "compact t" and "s \<inter> t = {}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5471
  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
  5472
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5473
  have *:"t \<inter> s = {}" using assms(3) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5474
  show ?thesis using separate_compact_closed[OF assms(2,1) *]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5475
    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
  5476
    by (auto simp add: dist_commute)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5477
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5478
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  5479
36439
a65320184de9 move intervals section heading
huffman
parents: 36438
diff changeset
  5480
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
  5481
  
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5482
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
  5483
  "{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
  5484
  "{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
  5485
  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
  5486
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5487
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
  5488
  "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
  5489
  "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
  5490
  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
  5491
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5492
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
  5493
 "({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
  5494
 "({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
  5495
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
  5496
  { 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
  5497
    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
  5498
    hence "a\<bullet>i < b\<bullet>i" by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5499
    hence False using as by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5500
  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
  5501
  { 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
  5502
    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
  5503
    { 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
  5504
      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
  5505
      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
  5506
        by (auto simp: inner_add_left) }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5507
    hence "{a <..< b} \<noteq> {}" using mem_interval(1)[of "?x" a b] by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5508
  ultimately show ?th1 by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5509
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
  5510
  { 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
  5511
    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
  5512
    hence "a\<bullet>i \<le> b\<bullet>i" by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5513
    hence False using as by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5514
  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
  5515
  { assume as:"\<forall>i\<in>Basis. \<not> (b\<bullet>i < a\<bullet>i)"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5516
    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
  5517
    { 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
  5518
      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
  5519
      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
  5520
        by (auto simp: inner_add_left) }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5521
    hence "{a .. b} \<noteq> {}" using mem_interval(2)[of "?x" a b] by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5522
  ultimately show ?th2 by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5523
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5524
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5525
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
  5526
  "{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
  5527
  "{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
  5528
  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
  5529
44584
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  5530
lemma interval_sing:
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  5531
  fixes a :: "'a::ordered_euclidean_space"
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  5532
  shows "{a .. a} = {a}" and "{a<..<a} = {}"
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  5533
  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
  5534
  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
  5535
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5536
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
  5537
 "(\<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
  5538
 "(\<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
  5539
 "(\<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
  5540
 "(\<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
  5541
  unfolding subset_eq[unfolded Ball_def] unfolding mem_interval
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  5542
  by (best intro: order_trans less_le_trans le_less_trans less_imp_le)+
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  5543
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  5544
lemma interval_open_subset_closed:
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  5545
  fixes a :: "'a::ordered_euclidean_space"
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  5546
  shows "{a<..<b} \<subseteq> {a .. b}"
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  5547
  unfolding subset_eq [unfolded Ball_def] mem_interval
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  5548
  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
  5549
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5550
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
  5551
 "{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
  5552
 "{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
  5553
 "{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
  5554
 "{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
  5555
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5556
  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
  5557
  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
  5558
  { 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
  5559
    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
  5560
    fix i :: 'a assume i:"i\<in>Basis"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5561
    (** 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
  5562
    { 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
  5563
      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
  5564
      { 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
  5565
        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
  5566
          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
  5567
          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
  5568
      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
  5569
      moreover
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5570
      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
  5571
        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
  5572
        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
  5573
        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
  5574
      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
  5575
    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
  5576
    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
  5577
    { 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
  5578
      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
  5579
      { 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
  5580
        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
  5581
          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
  5582
          by (auto simp add: as2) }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5583
      hence "?x\<in>{c<..<d}" unfolding mem_interval by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5584
      moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5585
      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
  5586
        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
  5587
        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
  5588
        by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5589
      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
  5590
    hence "b\<bullet>i \<ge> d\<bullet>i" by(rule ccontr)auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5591
    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
  5592
    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
  5593
  } 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
  5594
  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
  5595
    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
  5596
    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
  5597
    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
  5598
    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
  5599
    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
  5600
    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
  5601
    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
  5602
  { 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
  5603
    fix i :: 'a assume i:"i\<in>Basis"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5604
    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
  5605
    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
  5606
  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
  5607
    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
  5608
    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
  5609
qed
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5610
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5611
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
  5612
 "{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
  5613
  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
  5614
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5615
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
  5616
  "{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
  5617
  "{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
  5618
  "{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
  5619
  "{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
  5620
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
  5621
  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
  5622
  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
  5623
      (\<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
  5624
    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
  5625
  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
  5626
  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
  5627
  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
  5628
  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
  5629
  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
  5630
qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5631
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5632
(* Moved interval_open_subset_closed a bit upwards *)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5633
44250
9133bc634d9c simplify proofs of lemmas open_interval, closed_interval
huffman
parents: 44233
diff changeset
  5634
lemma open_interval[intro]:
9133bc634d9c simplify proofs of lemmas open_interval, closed_interval
huffman
parents: 44233
diff changeset
  5635
  fixes a b :: "'a::ordered_euclidean_space" shows "open {a<..<b}"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5636
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
  5637
  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
  5638
    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
  5639
      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
  5640
  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
  5641
    by (auto simp add: eucl_less [where 'a='a])
9133bc634d9c simplify proofs of lemmas open_interval, closed_interval
huffman
parents: 44233
diff changeset
  5642
  finally show "open {a<..<b}" .
9133bc634d9c simplify proofs of lemmas open_interval, closed_interval
huffman
parents: 44233
diff changeset
  5643
qed
9133bc634d9c simplify proofs of lemmas open_interval, closed_interval
huffman
parents: 44233
diff changeset
  5644
9133bc634d9c simplify proofs of lemmas open_interval, closed_interval
huffman
parents: 44233
diff changeset
  5645
lemma closed_interval[intro]:
9133bc634d9c simplify proofs of lemmas open_interval, closed_interval
huffman
parents: 44233
diff changeset
  5646
  fixes a b :: "'a::ordered_euclidean_space" shows "closed {a .. b}"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5647
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
  5648
  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
  5649
    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
  5650
      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
  5651
  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
  5652
    by (auto simp add: eucl_le [where 'a='a])
9133bc634d9c simplify proofs of lemmas open_interval, closed_interval
huffman
parents: 44233
diff changeset
  5653
  finally show "closed {a .. b}" .
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5654
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5655
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  5656
lemma interior_closed_interval [intro]:
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  5657
  fixes a b :: "'a::ordered_euclidean_space"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  5658
  shows "interior {a..b} = {a<..<b}" (is "?L = ?R")
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5659
proof(rule subset_antisym)
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  5660
  show "?R \<subseteq> ?L" using interval_open_subset_closed open_interval
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  5661
    by (rule interior_maximal)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5662
next
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  5663
  { fix x assume "x \<in> interior {a..b}"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  5664
    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
  5665
    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
  5666
    { 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
  5667
      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
  5668
           "dist (x + (e / 2) *\<^sub>R i) x < e"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5669
        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
  5670
        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
  5671
      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
  5672
                     "(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
  5673
        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
  5674
        and   e[THEN spec[where x="x + (e/2) *\<^sub>R i"]]
44584
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  5675
        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
  5676
      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
  5677
        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
  5678
    hence "x \<in> {a<..<b}" unfolding mem_interval by auto  }
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  5679
  thus "?L \<subseteq> ?R" ..
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5680
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5681
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5682
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
  5683
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
  5684
  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
  5685
  { 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
  5686
    { 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
  5687
      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
  5688
    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
  5689
    hence "norm x \<le> ?b" using norm_le_l1[of x] by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5690
  thus ?thesis unfolding interval and bounded_iff by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5691
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5692
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5693
lemma bounded_interval: fixes a :: "'a::ordered_euclidean_space" shows
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5694
 "bounded {a .. b} \<and> bounded {a<..<b}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5695
  using bounded_closed_interval[of a b]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5696
  using interval_open_subset_closed[of a b]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5697
  using bounded_subset[of "{a..b}" "{a<..<b}"]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5698
  by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5699
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
lemma not_interval_univ: fixes a :: "'a::ordered_euclidean_space" shows
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5701
 "({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
  5702
  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
  5703
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5704
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
  5705
  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
  5706
  by (auto simp: compact_eq_seq_compact_metric)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5707
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5708
lemma open_interval_midpoint: fixes a :: "'a::ordered_euclidean_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5709
  assumes "{a<..<b} \<noteq> {}" shows "((1/2) *\<^sub>R (a + b)) \<in> {a<..<b}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5710
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
  5711
  { 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
  5712
    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
  5713
      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
  5714
  thus ?thesis unfolding mem_interval by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5715
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5716
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5717
lemma open_closed_interval_convex: fixes x :: "'a::ordered_euclidean_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5718
  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
  5719
  shows "(e *\<^sub>R x + (1 - e) *\<^sub>R y) \<in> {a<..<b}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5720
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
  5721
  { 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
  5722
    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
  5723
    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
  5724
      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
  5725
      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
  5726
      using y unfolding mem_interval using i apply simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5727
      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
  5728
    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
  5729
    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
  5730
    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
  5731
    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
  5732
      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
  5733
      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
  5734
      using y unfolding mem_interval using i apply simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5735
      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
  5736
    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
  5737
    } 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
  5738
  thus ?thesis unfolding mem_interval by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5739
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5740
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5741
lemma closure_open_interval: fixes a :: "'a::ordered_euclidean_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5742
  assumes "{a<..<b} \<noteq> {}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5743
  shows "closure {a<..<b} = {a .. b}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5744
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
  5745
  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
  5746
  let ?c = "(1 / 2) *\<^sub>R (a + b)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5747
  { fix x assume as:"x \<in> {a .. b}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5748
    def f == "\<lambda>n::nat. x + (inverse (real n + 1)) *\<^sub>R (?c - x)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5749
    { 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
  5750
      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
  5751
      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
  5752
        x + (inverse (real n + 1)) *\<^sub>R (((1 / 2) *\<^sub>R (a + b)) - x)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5753
        by (auto simp add: algebra_simps)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5754
      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
  5755
      hence False using fn unfolding f_def using xc by auto  }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5756
    moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5757
    { assume "\<not> (f ---> x) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5758
      { fix e::real assume "e>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5759
        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
  5760
        then obtain N::nat where "inverse (real (N + 1)) < e" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5761
        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
  5762
        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
  5763
      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
  5764
        unfolding LIMSEQ_def by(auto simp add: dist_norm)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5765
      hence "(f ---> x) sequentially" unfolding f_def
44125
230a8665c919 mark some redundant theorems as legacy
huffman
parents: 44122
diff changeset
  5766
        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
  5767
        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
  5768
    ultimately have "x \<in> closure {a<..<b}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5769
      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
  5770
  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
  5771
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5772
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5773
lemma bounded_subset_open_interval_symmetric: fixes s::"('a::ordered_euclidean_space) set"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5774
  assumes "bounded s"  shows "\<exists>a. s \<subseteq> {-a<..<a}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5775
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5776
  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
  5777
  def a \<equiv> "(\<Sum>i\<in>Basis. (b + 1) *\<^sub>R i)::'a"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5778
  { 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
  5779
    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
  5780
    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
  5781
      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
  5782
  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
  5783
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5784
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5785
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
  5786
  fixes s :: "('a::ordered_euclidean_space) set"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5787
  shows "bounded s ==> (\<exists>a b. s \<subseteq> {a<..<b})"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5788
  by (auto dest!: bounded_subset_open_interval_symmetric)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5789
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5790
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
  5791
  fixes s :: "('a::ordered_euclidean_space) set"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5792
  assumes "bounded s" shows "\<exists>a. s \<subseteq> {-a .. a}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5793
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5794
  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
  5795
  thus ?thesis using interval_open_subset_closed[of "-a" a] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5796
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5797
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5798
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
  5799
  fixes s :: "('a::ordered_euclidean_space) set"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5800
  shows "bounded s ==> (\<exists>a b. s \<subseteq> {a .. b})"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5801
  using bounded_subset_closed_interval_symmetric[of s] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5802
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5803
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
  5804
  fixes a b :: "'a::ordered_euclidean_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5805
  shows "frontier {a .. b} = {a .. b} - {a<..<b}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5806
  unfolding frontier_def unfolding interior_closed_interval and closure_closed[OF closed_interval] ..
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5807
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5808
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
  5809
  fixes a b :: "'a::ordered_euclidean_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5810
  shows "frontier {a<..<b} = (if {a<..<b} = {} then {} else {a .. b} - {a<..<b})"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5811
proof(cases "{a<..<b} = {}")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5812
  case True thus ?thesis using frontier_empty by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5813
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5814
  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
  5815
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5816
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5817
lemma inter_interval_mixed_eq_empty: fixes a :: "'a::ordered_euclidean_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5818
  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
  5819
  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
  5820
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5821
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5822
(* Some stuff for half-infinite intervals too; FIXME: notation?  *)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5823
37673
f69f4b079275 generalize more lemmas from ordered_euclidean_space to euclidean_space
huffman
parents: 37649
diff changeset
  5824
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
  5825
  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
  5826
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
  5827
  { 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
  5828
    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
  5829
    { 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
  5830
      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
  5831
        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
  5832
      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
  5833
        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
  5834
    hence "x\<bullet>i \<le> b\<bullet>i" by(rule ccontr)auto  }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5835
  thus ?thesis unfolding closed_limpt unfolding islimpt_approachable by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5836
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5837
37673
f69f4b079275 generalize more lemmas from ordered_euclidean_space to euclidean_space
huffman
parents: 37649
diff changeset
  5838
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
  5839
  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
  5840
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
  5841
  { 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
  5842
    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
  5843
    { 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
  5844
      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
  5845
        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
  5846
      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
  5847
    hence "a\<bullet>i \<le> x\<bullet>i" by(rule ccontr)auto  }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5848
  thus ?thesis unfolding closed_limpt unfolding islimpt_approachable by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5849
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5850
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
  5851
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
  5852
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
  5853
  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
  5854
    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
  5855
  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
  5856
    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
  5857
  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
  5858
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
  5859
50881
ae630bab13da renamed countable_basis_space to second_countable_topology
hoelzl
parents: 50526
diff changeset
  5860
instance euclidean_space \<subseteq> second_countable_topology
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  5861
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
  5862
  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
  5863
  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
  5864
  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
  5865
  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
  5866
  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
  5867
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5868
  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
  5869
    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
  5870
  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
  5871
  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
  5872
  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
  5873
  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
  5874
    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
  5875
    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
  5876
      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
  5877
      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
  5878
      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
  5879
      done
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  5880
  qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  5881
  ultimately
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
  5882
  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
  5883
qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  5884
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  5885
instance ordered_euclidean_space \<subseteq> polish_space ..
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  5886
36439
a65320184de9 move intervals section heading
huffman
parents: 36438
diff changeset
  5887
text {* Intervals in general, including infinite and mixtures of open and closed. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5888
37732
6432bf0d7191 generalize type of is_interval to class euclidean_space
huffman
parents: 37680
diff changeset
  5889
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
  5890
  (\<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
  5891
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5892
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
  5893
  "is_interval {a<..<b}" (is ?th2) proof -
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5894
  show ?th1 ?th2  unfolding is_interval_def mem_interval Ball_def atLeastAtMost_iff
44584
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  5895
    by(meson order_trans le_less_trans less_le_trans less_trans)+ qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5896
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5897
lemma is_interval_empty:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5898
 "is_interval {}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5899
  unfolding is_interval_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5900
  by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5901
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5902
lemma is_interval_univ:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5903
 "is_interval UNIV"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5904
  unfolding is_interval_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5905
  by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5906
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  5907
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  5908
subsection {* Closure of halfspaces and hyperplanes *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5909
44219
f738e3200e24 generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
huffman
parents: 44216
diff changeset
  5910
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
  5911
  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
  5912
proof -
f738e3200e24 generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
huffman
parents: 44216
diff changeset
  5913
  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
  5914
    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
  5915
  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
  5916
    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
  5917
  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
  5918
    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
  5919
qed
f738e3200e24 generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
huffman
parents: 44216
diff changeset
  5920
f738e3200e24 generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
huffman
parents: 44216
diff changeset
  5921
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
  5922
  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
  5923
  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
  5924
  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
  5925
44213
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5926
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
  5927
  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
  5928
  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
  5929
  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
  5930
  shows "open {x. f x < g x}"
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5931
proof -
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5932
  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
  5933
    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
  5934
    by (rule isCont_open_vimage)
44213
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5935
  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
  5936
    by auto
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5937
  finally show ?thesis .
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5938
qed
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5939
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5940
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
  5941
  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
  5942
  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
  5943
  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
  5944
  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
  5945
proof -
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5946
  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
  5947
    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
  5948
    by (rule isCont_closed_vimage)
44213
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5949
  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
  5950
    by auto
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5951
  finally show ?thesis .
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5952
qed
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5953
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5954
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
  5955
  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
  5956
  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
  5957
  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
  5958
  shows "closed {x. f x = g x}"
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5959
proof -
44216
903bfe95fece generalized lemma closed_Collect_eq
huffman
parents: 44213
diff changeset
  5960
  have "open {(x::'b, y::'b). x \<noteq> y}"
903bfe95fece generalized lemma closed_Collect_eq
huffman
parents: 44213
diff changeset
  5961
    unfolding open_prod_def by (auto dest!: hausdorff)
903bfe95fece generalized lemma closed_Collect_eq
huffman
parents: 44213
diff changeset
  5962
  hence "closed {(x::'b, y::'b). x = y}"
903bfe95fece generalized lemma closed_Collect_eq
huffman
parents: 44213
diff changeset
  5963
    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
  5964
  with isCont_Pair [OF f g]
44216
903bfe95fece generalized lemma closed_Collect_eq
huffman
parents: 44213
diff changeset
  5965
  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
  5966
    by (rule isCont_closed_vimage)
44216
903bfe95fece generalized lemma closed_Collect_eq
huffman
parents: 44213
diff changeset
  5967
  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
  5968
  finally show ?thesis .
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5969
qed
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5970
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5971
lemma continuous_at_inner: "continuous (at x) (inner a)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5972
  unfolding continuous_at by (intro tendsto_intros)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5973
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5974
lemma closed_halfspace_le: "closed {x. inner a x \<le> b}"
44233
aa74ce315bae add simp rules for isCont
huffman
parents: 44219
diff changeset
  5975
  by (simp add: closed_Collect_le)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5976
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5977
lemma closed_halfspace_ge: "closed {x. inner a x \<ge> b}"
44233
aa74ce315bae add simp rules for isCont
huffman
parents: 44219
diff changeset
  5978
  by (simp add: closed_Collect_le)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5979
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5980
lemma closed_hyperplane: "closed {x. inner a x = b}"
44233
aa74ce315bae add simp rules for isCont
huffman
parents: 44219
diff changeset
  5981
  by (simp add: closed_Collect_eq)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5982
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5983
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
  5984
  shows "closed {x::'a::euclidean_space. x\<bullet>i \<le> a}"
44233
aa74ce315bae add simp rules for isCont
huffman
parents: 44219
diff changeset
  5985
  by (simp add: closed_Collect_le)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5986
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5987
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
  5988
  shows "closed {x::'a::euclidean_space. x\<bullet>i \<ge> a}"
44233
aa74ce315bae add simp rules for isCont
huffman
parents: 44219
diff changeset
  5989
  by (simp add: closed_Collect_le)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5990
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  5991
text {* Openness of halfspaces. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5992
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5993
lemma open_halfspace_lt: "open {x. inner a x < b}"
44233
aa74ce315bae add simp rules for isCont
huffman
parents: 44219
diff changeset
  5994
  by (simp add: open_Collect_less)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5995
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5996
lemma open_halfspace_gt: "open {x. inner a x > b}"
44233
aa74ce315bae add simp rules for isCont
huffman
parents: 44219
diff changeset
  5997
  by (simp add: open_Collect_less)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5998
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5999
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
  6000
  shows "open {x::'a::euclidean_space. x\<bullet>i < a}"
44233
aa74ce315bae add simp rules for isCont
huffman
parents: 44219
diff changeset
  6001
  by (simp add: open_Collect_less)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6002
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6003
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
  6004
  shows "open {x::'a::euclidean_space. x\<bullet>i > a}"
44233
aa74ce315bae add simp rules for isCont
huffman
parents: 44219
diff changeset
  6005
  by (simp add: open_Collect_less)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6006
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  6007
text{* Instantiation for intervals on @{text ordered_euclidean_space} *}
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  6008
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  6009
lemma eucl_lessThan_eq_halfspaces:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  6010
  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
  6011
  shows "{..<a} = (\<Inter>i\<in>Basis. {x. x \<bullet> i < a \<bullet> i})"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  6012
 by (auto simp: eucl_less[where 'a='a])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  6013
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  6014
lemma eucl_greaterThan_eq_halfspaces:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  6015
  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
  6016
  shows "{a<..} = (\<Inter>i\<in>Basis. {x. a \<bullet> i < x \<bullet> i})"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  6017
 by (auto simp: eucl_less[where 'a='a])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  6018
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  6019
lemma eucl_atMost_eq_halfspaces:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  6020
  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
  6021
  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
  6022
 by (auto simp: eucl_le[where 'a='a])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  6023
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  6024
lemma eucl_atLeast_eq_halfspaces:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  6025
  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
  6026
  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
  6027
 by (auto simp: eucl_le[where 'a='a])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  6028
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  6029
lemma open_eucl_lessThan[simp, intro]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  6030
  fixes a :: "'a\<Colon>ordered_euclidean_space"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  6031
  shows "open {..< a}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  6032
  by (auto simp: eucl_lessThan_eq_halfspaces open_halfspace_component_lt)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  6033
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  6034
lemma open_eucl_greaterThan[simp, intro]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  6035
  fixes a :: "'a\<Colon>ordered_euclidean_space"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  6036
  shows "open {a <..}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  6037
  by (auto simp: eucl_greaterThan_eq_halfspaces open_halfspace_component_gt)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  6038
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  6039
lemma closed_eucl_atMost[simp, intro]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  6040
  fixes a :: "'a\<Colon>ordered_euclidean_space"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  6041
  shows "closed {.. a}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  6042
  unfolding eucl_atMost_eq_halfspaces
44233
aa74ce315bae add simp rules for isCont
huffman
parents: 44219
diff changeset
  6043
  by (simp add: closed_INT closed_Collect_le)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  6044
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  6045
lemma closed_eucl_atLeast[simp, intro]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  6046
  fixes a :: "'a\<Colon>ordered_euclidean_space"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  6047
  shows "closed {a ..}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  6048
  unfolding eucl_atLeast_eq_halfspaces
44233
aa74ce315bae add simp rules for isCont
huffman
parents: 44219
diff changeset
  6049
  by (simp add: closed_INT closed_Collect_le)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  6050
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  6051
text {* This gives a simple derivation of limit component bounds. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6052
37673
f69f4b079275 generalize more lemmas from ordered_euclidean_space to euclidean_space
huffman
parents: 37649
diff changeset
  6053
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
  6054
  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
  6055
  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
  6056
  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
  6057
37673
f69f4b079275 generalize more lemmas from ordered_euclidean_space to euclidean_space
huffman
parents: 37649
diff changeset
  6058
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
  6059
  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
  6060
  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
  6061
  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
  6062
37673
f69f4b079275 generalize more lemmas from ordered_euclidean_space to euclidean_space
huffman
parents: 37649
diff changeset
  6063
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
  6064
  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
  6065
  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
  6066
  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
  6067
  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
  6068
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6069
text{* Limits relative to a union.                                               *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6070
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6071
lemma eventually_within_Un:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6072
  "eventually P (net within (s \<union> t)) \<longleftrightarrow>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6073
    eventually P (net within s) \<and> eventually P (net within t)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6074
  unfolding Limits.eventually_within
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6075
  by (auto elim!: eventually_rev_mp)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6076
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6077
lemma Lim_within_union:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6078
 "(f ---> l) (net within (s \<union> t)) \<longleftrightarrow>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6079
  (f ---> l) (net within s) \<and> (f ---> l) (net within t)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6080
  unfolding tendsto_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6081
  by (auto simp add: eventually_within_Un)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6082
36442
b96d9dc6acca generalize more continuity lemmas
huffman
parents: 36441
diff changeset
  6083
lemma Lim_topological:
b96d9dc6acca generalize more continuity lemmas
huffman
parents: 36441
diff changeset
  6084
 "(f ---> l) net \<longleftrightarrow>
b96d9dc6acca generalize more continuity lemmas
huffman
parents: 36441
diff changeset
  6085
        trivial_limit net \<or>
b96d9dc6acca generalize more continuity lemmas
huffman
parents: 36441
diff changeset
  6086
        (\<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
  6087
  unfolding tendsto_def trivial_limit_eq by auto
b96d9dc6acca generalize more continuity lemmas
huffman
parents: 36441
diff changeset
  6088
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6089
lemma continuous_on_union:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6090
  assumes "closed s" "closed t" "continuous_on s f" "continuous_on t f"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6091
  shows "continuous_on (s \<union> t) f"
36442
b96d9dc6acca generalize more continuity lemmas
huffman
parents: 36441
diff changeset
  6092
  using assms unfolding continuous_on Lim_within_union
b96d9dc6acca generalize more continuity lemmas
huffman
parents: 36441
diff changeset
  6093
  unfolding Lim_topological trivial_limit_within closed_limpt by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6094
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6095
lemma continuous_on_cases:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6096
  assumes "closed s" "closed t" "continuous_on s f" "continuous_on t g"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6097
          "\<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
  6098
  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
  6099
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6100
  let ?h = "(\<lambda>x. if P x then f x else g x)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6101
  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
  6102
  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
  6103
  moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6104
  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
  6105
  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
  6106
  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
  6107
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6108
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6109
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6110
text{* Some more convenient intermediate-value theorem formulations.             *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6111
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6112
lemma connected_ivt_hyperplane:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6113
  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
  6114
  shows "\<exists>z \<in> s. inner a z = b"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6115
proof(rule ccontr)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6116
  assume as:"\<not> (\<exists>z\<in>s. inner a z = b)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6117
  let ?A = "{x. inner a x < b}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6118
  let ?B = "{x. inner a x > b}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6119
  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
  6120
  moreover have "?A \<inter> ?B = {}" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6121
  moreover have "s \<subseteq> ?A \<union> ?B" using as by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6122
  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
  6123
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6124
37673
f69f4b079275 generalize more lemmas from ordered_euclidean_space to euclidean_space
huffman
parents: 37649
diff changeset
  6125
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
  6126
 "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
  6127
  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
  6128
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  6129
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  6130
subsection {* Homeomorphisms *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6131
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6132
definition "homeomorphism s t f g \<equiv>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6133
     (\<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
  6134
     (\<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
  6135
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6136
definition
50898
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  6137
  homeomorphic :: "'a::topological_space set \<Rightarrow> 'b::topological_space set \<Rightarrow> bool"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6138
    (infixr "homeomorphic" 60) where
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6139
  homeomorphic_def: "s homeomorphic t \<equiv> (\<exists>f g. homeomorphism s t f g)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6140
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6141
lemma homeomorphic_refl: "s homeomorphic s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6142
  unfolding homeomorphic_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6143
  unfolding homeomorphism_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6144
  using continuous_on_id
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6145
  apply(rule_tac x = "(\<lambda>x. x)" in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6146
  apply(rule_tac x = "(\<lambda>x. x)" in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6147
  by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6148
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6149
lemma homeomorphic_sym:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6150
 "s homeomorphic t \<longleftrightarrow> t homeomorphic s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6151
unfolding homeomorphic_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6152
unfolding homeomorphism_def
33324
51eb2ffa2189 Tidied up some very ugly proofs
paulson
parents: 33270
diff changeset
  6153
by blast 
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6154
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6155
lemma homeomorphic_trans:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6156
  assumes "s homeomorphic t" "t homeomorphic u" shows "s homeomorphic u"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6157
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6158
  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
  6159
    using assms(1) unfolding homeomorphic_def homeomorphism_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6160
  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
  6161
    using assms(2) unfolding homeomorphic_def homeomorphism_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6162
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6163
  { 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
  6164
  moreover have "(f2 \<circ> f1) ` s = u" using fg1(2) fg2(2) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6165
  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
  6166
  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
  6167
  moreover have "(g1 \<circ> g2) ` u = s" using fg1(5) fg2(5) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6168
  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
  6169
  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
  6170
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6171
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6172
lemma homeomorphic_minimal:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6173
 "s homeomorphic t \<longleftrightarrow>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6174
    (\<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
  6175
           (\<forall>y\<in>t. g(y) \<in> s \<and> (f(g(y)) = y)) \<and>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6176
           continuous_on s f \<and> continuous_on t g)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6177
unfolding homeomorphic_def homeomorphism_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6178
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
  6179
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
  6180
unfolding image_iff
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6181
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
  6182
apply auto apply(rule_tac x="g x" in bexI) apply auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6183
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
  6184
apply auto apply(rule_tac x="f x" in bexI) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6185
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  6186
text {* Relatively weak hypotheses if a set is compact. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6187
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6188
lemma homeomorphism_compact:
50898
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  6189
  fixes f :: "'a::topological_space \<Rightarrow> 'b::t2_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6190
  assumes "compact s" "continuous_on s f"  "f ` s = t"  "inj_on f s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6191
  shows "\<exists>g. homeomorphism s t f g"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6192
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6193
  def g \<equiv> "\<lambda>x. SOME y. y\<in>s \<and> f y = x"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6194
  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
  6195
  { fix y assume "y\<in>t"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6196
    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
  6197
    hence "g (f x) = x" using g by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6198
    hence "f (g y) = y" unfolding x(1)[THEN sym] by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6199
  hence g':"\<forall>x\<in>t. f (g x) = x" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6200
  moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6201
  { fix x
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6202
    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
  6203
    moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6204
    { assume "x\<in>g ` t"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6205
      then obtain y where y:"y\<in>t" "g y = x" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6206
      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
  6207
      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
  6208
    ultimately have "x\<in>s \<longleftrightarrow> x \<in> g ` t" ..  }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6209
  hence "g ` t = s" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6210
  ultimately
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6211
  show ?thesis unfolding homeomorphism_def homeomorphic_def
44647
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  6212
    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
  6213
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6214
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6215
lemma homeomorphic_compact:
50898
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  6216
  fixes f :: "'a::topological_space \<Rightarrow> 'b::t2_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6217
  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
  6218
          \<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
  6219
  unfolding homeomorphic_def by (metis homeomorphism_compact)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6220
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6221
text{* Preservation of topological properties.                                   *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6222
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6223
lemma homeomorphic_compactness:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6224
 "s homeomorphic t ==> (compact s \<longleftrightarrow> compact t)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6225
unfolding homeomorphic_def homeomorphism_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6226
by (metis compact_continuous_image)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6227
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6228
text{* Results on translation, scaling etc.                                      *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6229
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6230
lemma homeomorphic_scaling:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6231
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6232
  assumes "c \<noteq> 0"  shows "s homeomorphic ((\<lambda>x. c *\<^sub>R x) ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6233
  unfolding homeomorphic_minimal
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6234
  apply(rule_tac x="\<lambda>x. c *\<^sub>R x" in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6235
  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
  6236
  using assms by (auto simp add: continuous_on_intros)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6237
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6238
lemma homeomorphic_translation:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6239
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6240
  shows "s homeomorphic ((\<lambda>x. a + x) ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6241
  unfolding homeomorphic_minimal
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6242
  apply(rule_tac x="\<lambda>x. a + x" in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6243
  apply(rule_tac x="\<lambda>x. -a + x" in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6244
  using continuous_on_add[OF continuous_on_const continuous_on_id] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6245
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6246
lemma homeomorphic_affinity:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6247
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6248
  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
  6249
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6250
  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
  6251
  show ?thesis
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6252
    using homeomorphic_trans
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6253
    using homeomorphic_scaling[OF assms, of s]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6254
    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
  6255
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6256
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6257
lemma homeomorphic_balls:
50898
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  6258
  fixes a b ::"'a::real_normed_vector"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6259
  assumes "0 < d"  "0 < e"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6260
  shows "(ball a d) homeomorphic  (ball b e)" (is ?th)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6261
        "(cball a d) homeomorphic (cball b e)" (is ?cth)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6262
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6263
  show ?th unfolding homeomorphic_minimal
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6264
    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
  6265
    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
  6266
    using assms apply (auto simp add: dist_commute)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6267
    unfolding dist_norm
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6268
    apply (auto simp add: pos_divide_less_eq mult_strict_left_mono)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6269
    unfolding continuous_on
36659
f794e92784aa adapt to removed premise on tendsto lemma (cf. 88f0125c3bd2)
huffman
parents: 36623
diff changeset
  6270
    by (intro ballI tendsto_intros, simp)+
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6271
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6272
  show ?cth unfolding homeomorphic_minimal
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6273
    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
  6274
    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
  6275
    using assms apply (auto simp add: dist_commute)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6276
    unfolding dist_norm
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6277
    apply (auto simp add: pos_divide_le_eq)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6278
    unfolding continuous_on
36659
f794e92784aa adapt to removed premise on tendsto lemma (cf. 88f0125c3bd2)
huffman
parents: 36623
diff changeset
  6279
    by (intro ballI tendsto_intros, simp)+
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6280
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6281
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6282
text{* "Isometry" (up to constant bounds) of injective linear map etc.           *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6283
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6284
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
  6285
  fixes x :: "nat \<Rightarrow> 'a::euclidean_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6286
  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
  6287
  shows "Cauchy x"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6288
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6289
  interpret f: bounded_linear f by fact
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6290
  { fix d::real assume "d>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6291
    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
  6292
      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
  6293
    { 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
  6294
      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
  6295
        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
  6296
        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
  6297
      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
  6298
        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
  6299
      finally have "norm (x n - x N) < d" using `e>0` by simp }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6300
    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
  6301
  thus ?thesis unfolding cauchy and dist_norm by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6302
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6303
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6304
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
  6305
  fixes f :: "'a::euclidean_space => 'b::euclidean_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6306
  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
  6307
  shows "complete(f ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6308
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6309
  { 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
  6310
    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
  6311
      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
  6312
    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
  6313
    hence "f \<circ> x = g" unfolding fun_eq_iff by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6314
    then obtain l where "l\<in>s" and l:"(x ---> l) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6315
      using cs[unfolded complete_def, THEN spec[where x="x"]]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6316
      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
  6317
    hence "\<exists>l\<in>f ` s. (g ---> l) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6318
      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
  6319
      unfolding `f \<circ> x = g` by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6320
  thus ?thesis unfolding complete_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6321
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6322
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6323
lemma dist_0_norm:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6324
  fixes x :: "'a::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6325
  shows "dist 0 x = norm x"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6326
unfolding dist_norm by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6327
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  6328
lemma injective_imp_isometric: fixes f::"'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6329
  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
  6330
  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
  6331
proof(cases "s \<subseteq> {0::'a}")
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6332
  case True
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6333
  { fix x assume "x \<in> s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6334
    hence "x = 0" using True by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6335
    hence "norm x \<le> norm (f x)" by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6336
  thus ?thesis by(auto intro!: exI[where x=1])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6337
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6338
  interpret f: bounded_linear f by fact
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6339
  case False
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6340
  then obtain a where a:"a\<noteq>0" "a\<in>s" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6341
  from False have "s \<noteq> {}" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6342
  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
  6343
  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
  6344
  let ?S'' = "{x::'a. norm x = norm a}"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6345
36362
06475a1547cb fix lots of looping simp calls and other warnings
huffman
parents: 36360
diff changeset
  6346
  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
  6347
  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
  6348
  moreover have "?S' = s \<inter> ?S''" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6349
  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
  6350
  moreover have *:"f ` ?S' = ?S" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6351
  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
  6352
  hence "closed ?S" using compact_imp_closed by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6353
  moreover have "?S \<noteq> {}" using a by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6354
  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
  6355
  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
  6356
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6357
  let ?e = "norm (f b) / norm b"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6358
  have "norm b > 0" using ba and a and norm_ge_zero by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6359
  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
  6360
  ultimately have "0 < norm (f b) / norm b" by(simp only: divide_pos_pos)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6361
  moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6362
  { fix x assume "x\<in>s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6363
    hence "norm (f b) / norm b * norm x \<le> norm (f x)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6364
    proof(cases "x=0")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6365
      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
  6366
    next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6367
      case False
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6368
      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
  6369
      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
  6370
      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
  6371
      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
  6372
        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
  6373
        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
  6374
    qed }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6375
  ultimately
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6376
  show ?thesis by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6377
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6378
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6379
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
  6380
  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6381
  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
  6382
  shows "closed(f ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6383
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6384
  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
  6385
  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
  6386
    unfolding complete_eq_closed[THEN sym] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6387
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6388
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  6389
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  6390
subsection {* Some properties of a canonical subspace *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6391
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6392
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
  6393
  "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
  6394
  unfolding subspace_def by (auto simp: inner_add_left)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6395
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6396
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
  6397
 "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
  6398
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
  6399
  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
  6400
  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
  6401
    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
  6402
  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
  6403
    by auto
d366fa5551ef declare euclidean_simps [simp] at the point they are proved;
huffman
parents: 44365
diff changeset
  6404
  finally show "closed ?A" .
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6405
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6406
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
  6407
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
  6408
  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
  6409
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
  6410
  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
  6411
  have "d \<subseteq> ?A" using d by (auto simp: inner_Basis)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6412
  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
  6413
  { 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
  6414
    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
  6415
    from this d have "x \<in> span d"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6416
    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
  6417
      case empty hence "x=0" apply(rule_tac euclidean_eqI) by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6418
      thus ?case using subspace_0[OF subspace_span[of "{}"]] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6419
    next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6420
      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
  6421
      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
  6422
      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
  6423
      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
  6424
      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
  6425
      { 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
  6426
        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
  6427
          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
  6428
      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
  6429
      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
  6430
        using span_mono[of F "insert k F"] using assms by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6431
      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
  6432
      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
  6433
      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
  6434
        using span_mul by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6435
      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
  6436
      have "y + (x\<bullet>k) *\<^sub>R k \<in> span (insert k F)"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6437
        using span_add by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6438
      thus ?case using y by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6439
    qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6440
  }
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
  6441
  hence "?A \<subseteq> span d" by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6442
  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
  6443
  { 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
  6444
  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
  6445
  moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6446
  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
  6447
  ultimately show ?thesis using dim_unique[of d ?A] by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6448
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6449
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6450
text{* Hence closure and completeness of all subspaces.                          *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6451
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
  6452
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
  6453
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
  6454
  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
  6455
  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
  6456
  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
  6457
    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
  6458
  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
  6459
    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
  6460
  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
  6461
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
  6462
  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
  6463
qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6464
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  6465
lemma closed_subspace: fixes s::"('a::euclidean_space) set"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6466
  assumes "subspace s" shows "closed s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6467
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
  6468
  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
  6469
  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
  6470
  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
  6471
  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
  6472
      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
  6473
    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
  6474
    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
  6475
  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
  6476
  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
  6477
  { 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
  6478
  moreover have "closed ?t" using closed_substandard .
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6479
  moreover have "subspace ?t" using subspace_substandard .
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6480
  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
  6481
    unfolding f(2) using f(1) unfolding linear_conv_bounded_linear by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6482
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6483
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6484
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
  6485
  fixes s :: "('a::euclidean_space) set" shows "subspace s ==> complete s"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6486
  using complete_eq_closed closed_subspace
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6487
  by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6488
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6489
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
  6490
  fixes s :: "('a::euclidean_space) set"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6491
  shows "dim(closure s) = dim s" (is "?dc = ?d")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6492
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6493
  have "?dc \<le> ?d" using closure_minimal[OF span_inc, of s]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6494
    using closed_subspace[OF subspace_span, of s]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6495
    using dim_subset[of "closure s" "span s"] unfolding dim_span by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6496
  thus ?thesis using dim_subset[OF closure_subset, of s] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6497
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6498
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  6499
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  6500
subsection {* Affine transformations of intervals *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6501
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6502
lemma real_affinity_le:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34999
diff changeset
  6503
 "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
  6504
  by (simp add: field_simps inverse_eq_divide)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6505
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6506
lemma real_le_affinity:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34999
diff changeset
  6507
 "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
  6508
  by (simp add: field_simps inverse_eq_divide)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6509
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6510
lemma real_affinity_lt:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34999
diff changeset
  6511
 "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
  6512
  by (simp add: field_simps inverse_eq_divide)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6513
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6514
lemma real_lt_affinity:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34999
diff changeset
  6515
 "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
  6516
  by (simp add: field_simps inverse_eq_divide)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6517
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6518
lemma real_affinity_eq:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34999
diff changeset
  6519
 "(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
  6520
  by (simp add: field_simps inverse_eq_divide)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6521
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6522
lemma real_eq_affinity:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34999
diff changeset
  6523
 "(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
  6524
  by (simp add: field_simps inverse_eq_divide)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6525
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6526
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
  6527
  fixes a b c :: "'a::ordered_euclidean_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6528
  shows "(\<lambda>x. m *\<^sub>R x + c) ` {a .. b} =
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6529
            (if {a .. b} = {} then {}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6530
            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
  6531
            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
  6532
proof(cases "m=0")  
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6533
  { 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
  6534
    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
  6535
      apply(subst euclidean_eq_iff) by (auto intro: order_antisym) }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6536
  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
  6537
  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
  6538
  ultimately show ?thesis by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6539
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6540
  case False
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6541
  { fix y assume "a \<le> y" "y \<le> b" "m > 0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6542
    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
  6543
      unfolding eucl_le[where 'a='a] by (auto simp: inner_simps)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6544
  } moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6545
  { fix y assume "a \<le> y" "y \<le> b" "m < 0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6546
    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
  6547
      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
  6548
  } moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6549
  { 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
  6550
    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
  6551
      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
  6552
      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
  6553
      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
  6554
  } moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6555
  { 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
  6556
    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
  6557
      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
  6558
      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
  6559
      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
  6560
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6561
  ultimately show ?thesis using False by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6562
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6563
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  6564
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
  6565
  (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
  6566
  using image_affinity_interval[of m 0 a b] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6567
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  6568
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  6569
subsection {* Banach fixed point theorem (not really topological...) *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6570
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6571
lemma banach_fix:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6572
  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
  6573
          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
  6574
  shows "\<exists>! x\<in>s. (f x = x)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6575
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6576
  have "1 - c > 0" using c by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6577
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6578
  from s(2) obtain z0 where "z0 \<in> s" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6579
  def z \<equiv> "\<lambda>n. (f ^^ n) z0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6580
  { fix n::nat
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6581
    have "z n \<in> s" unfolding z_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6582
    proof(induct n) case 0 thus ?case using `z0 \<in>s` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6583
    next case Suc thus ?case using f by auto qed }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6584
  note z_in_s = this
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6585
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6586
  def d \<equiv> "dist (z 0) (z 1)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6587
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6588
  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
  6589
  { fix n::nat
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6590
    have "dist (z n) (z (Suc n)) \<le> (c ^ n) * d"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6591
    proof(induct n)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6592
      case 0 thus ?case unfolding d_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6593
    next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6594
      case (Suc m)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6595
      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
  6596
        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
  6597
      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
  6598
        unfolding fzn and mult_le_cancel_left by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6599
    qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6600
  } note cf_z = this
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6601
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6602
  { fix n m::nat
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6603
    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
  6604
    proof(induct n)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6605
      case 0 show ?case by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6606
    next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6607
      case (Suc k)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6608
      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
  6609
        using dist_triangle and c by(auto simp add: dist_triangle)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6610
      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
  6611
        using cf_z[of "m + k"] and c by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6612
      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
  6613
        using Suc by (auto simp add: field_simps)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6614
      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
  6615
        unfolding power_add by (auto simp add: field_simps)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6616
      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
  6617
        using c by (auto simp add: field_simps)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6618
      finally show ?case by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6619
    qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6620
  } note cf_z2 = this
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6621
  { fix e::real assume "e>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6622
    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
  6623
    proof(cases "d = 0")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6624
      case True
41863
e5104b436ea1 removed dependency on Dense_Linear_Order
boehmes
parents: 41413
diff changeset
  6625
      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
  6626
        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
  6627
      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
  6628
        by (simp add: *)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6629
      thus ?thesis using `e>0` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6630
    next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6631
      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
  6632
        by (metis False d_def less_le)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6633
      hence "0 < e * (1 - c) / d" using `e>0` and `1-c>0`
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6634
        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
  6635
      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
  6636
      { 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
  6637
        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
  6638
        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
  6639
        hence **:"d * (1 - c ^ (m - n)) / (1 - c) > 0"
36778
739a9379e29b avoid using real-specific versions of generic lemmas
huffman
parents: 36669
diff changeset
  6640
          using mult_pos_pos[OF `d>0`, of "1 - c ^ (m - n)"]
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6641
          using divide_pos_pos[of "d * (1 - c ^ (m - n))" "1 - c"]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6642
          using `0 < 1 - c` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6643
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6644
        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
  6645
          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
  6646
          by (auto simp add: mult_commute dist_commute)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6647
        also have "\<dots> \<le> c ^ N * d * (1 - c ^ (m - n)) / (1 - c)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6648
          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
  6649
          unfolding mult_assoc by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6650
        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
  6651
          using mult_strict_right_mono[OF N **] unfolding mult_assoc by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6652
        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
  6653
        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
  6654
        finally have  "dist (z m) (z n) < e" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6655
      } note * = this
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6656
      { fix m n::nat assume as:"N\<le>m" "N\<le>n"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6657
        hence "dist (z n) (z m) < e"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6658
        proof(cases "n = m")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6659
          case True thus ?thesis using `e>0` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6660
        next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6661
          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
  6662
        qed }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6663
      thus ?thesis by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6664
    qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6665
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6666
  hence "Cauchy z" unfolding cauchy_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6667
  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
  6668
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6669
  def e \<equiv> "dist (f x) x"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6670
  have "e = 0" proof(rule ccontr)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6671
    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
  6672
      by (metis dist_eq_0_iff dist_nz e_def)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6673
    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
  6674
      using x[unfolded LIMSEQ_def, THEN spec[where x="e/2"]] by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6675
    hence N':"dist (z N) x < e / 2" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6676
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6677
    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
  6678
      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
  6679
      by (metis dist_eq_0_iff dist_nz order_less_asym less_le)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6680
    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
  6681
      using z_in_s[of N] `x\<in>s` using c by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6682
    also have "\<dots> < e / 2" using N' and c using * by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6683
    finally show False unfolding fzn
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6684
      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
  6685
      unfolding e_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6686
  qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6687
  hence "f x = x" unfolding e_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6688
  moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6689
  { fix y assume "f y = y" "y\<in>s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6690
    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
  6691
      using `x\<in>s` and `f x = x` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6692
    hence "dist x y = 0" unfolding mult_le_cancel_right1
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6693
      using c and zero_le_dist[of x y] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6694
    hence "y = x" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6695
  }
34999
5312d2ffee3b Changed 'bounded unique existential quantifiers' from a constant to syntax translation.
hoelzl
parents: 34964
diff changeset
  6696
  ultimately show ?thesis using `x\<in>s` by blast+
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6697
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6698
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  6699
subsection {* Edelstein fixed point theorem *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6700
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6701
lemma edelstein_fix:
50970
3e5b67f85bf9 generalized theorem edelstein_fix to class metric_space
huffman
parents: 50955
diff changeset
  6702
  fixes s :: "'a::metric_space set"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6703
  assumes s:"compact s" "s \<noteq> {}" and gs:"(g ` s) \<subseteq> s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6704
      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
  6705
  shows "\<exists>! x\<in>s. g x = x"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6706
proof(cases "\<exists>x\<in>s. g x \<noteq> x")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6707
  obtain x where "x\<in>s" using s(2) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6708
  case False hence g:"\<forall>x\<in>s. g x = x" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6709
  { fix y assume "y\<in>s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6710
    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
  6711
      unfolding g[THEN bspec[where x=x], OF `x\<in>s`]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6712
      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
  6713
  thus ?thesis using `x\<in>s` and g by blast+
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6714
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6715
  case True
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6716
  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
  6717
  { fix x y assume "x \<in> s" "y \<in> s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6718
    hence "dist (g x) (g y) \<le> dist x y"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6719
      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
  6720
  def y \<equiv> "g x"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6721
  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
  6722
  def f \<equiv> "\<lambda>n. g ^^ n"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6723
  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
  6724
  have [simp]:"\<And>z. f 0 z = z" unfolding f_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6725
  { fix n::nat and z assume "z\<in>s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6726
    have "f n z \<in> s" unfolding f_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6727
    proof(induct n)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6728
      case 0 thus ?case using `z\<in>s` by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6729
    next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6730
      case (Suc n) thus ?case using gs[unfolded image_subset_iff] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6731
    qed } note fs = this
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6732
  { fix m n ::nat assume "m\<le>n"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6733
    fix w z assume "w\<in>s" "z\<in>s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6734
    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
  6735
    proof(induct n)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6736
      case 0 thus ?case by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6737
    next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6738
      case (Suc n)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6739
      thus ?case proof(cases "m\<le>n")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6740
        case True thus ?thesis using Suc(1)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6741
          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
  6742
      next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6743
        case False hence mn:"m = Suc n" using Suc(2) by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6744
        show ?thesis unfolding mn  by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6745
      qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6746
    qed } note distf = this
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6747
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6748
  def h \<equiv> "\<lambda>n. (f n x, f n y)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6749
  let ?s2 = "s \<times> s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6750
  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
  6751
    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
  6752
    using fs[OF `x\<in>s`] and fs[OF `y\<in>s`] by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6753
  def a \<equiv> "fst l" def b \<equiv> "snd l"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6754
  have lab:"l = (a, b)" unfolding a_def b_def by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6755
  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
  6756
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6757
  have lima:"((fst \<circ> (h \<circ> r)) ---> a) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6758
   and limb:"((snd \<circ> (h \<circ> r)) ---> b) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6759
    using lr
44167
e81d676d598e avoid duplicate rule warnings
huffman
parents: 44139
diff changeset
  6760
    unfolding o_def a_def b_def by (rule tendsto_intros)+
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6761
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6762
  { fix n::nat
50970
3e5b67f85bf9 generalized theorem edelstein_fix to class metric_space
huffman
parents: 50955
diff changeset
  6763
    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
  6764
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6765
    { assume as:"dist a b > dist (f n x) (f n y)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6766
      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
  6767
        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
  6768
        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
  6769
      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
  6770
        apply -
3e5b67f85bf9 generalized theorem edelstein_fix to class metric_space
huffman
parents: 50955
diff changeset
  6771
        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
  6772
        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
  6773
        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
  6774
        apply (erule le_less_trans [rotated])
3e5b67f85bf9 generalized theorem edelstein_fix to class metric_space
huffman
parents: 50955
diff changeset
  6775
        apply (erule thin_rl)
3e5b67f85bf9 generalized theorem edelstein_fix to class metric_space
huffman
parents: 50955
diff changeset
  6776
        apply (rule abs_leI)
3e5b67f85bf9 generalized theorem edelstein_fix to class metric_space
huffman
parents: 50955
diff changeset
  6777
        apply (simp add: diff_le_iff add_assoc)
3e5b67f85bf9 generalized theorem edelstein_fix to class metric_space
huffman
parents: 50955
diff changeset
  6778
        apply (rule order_trans [OF dist_triangle add_left_mono])
3e5b67f85bf9 generalized theorem edelstein_fix to class metric_space
huffman
parents: 50955
diff changeset
  6779
        apply (subst add_commute, rule dist_triangle2)
3e5b67f85bf9 generalized theorem edelstein_fix to class metric_space
huffman
parents: 50955
diff changeset
  6780
        apply (simp add: diff_le_iff add_assoc)
3e5b67f85bf9 generalized theorem edelstein_fix to class metric_space
huffman
parents: 50955
diff changeset
  6781
        apply (rule order_trans [OF dist_triangle3 add_left_mono])
3e5b67f85bf9 generalized theorem edelstein_fix to class metric_space
huffman
parents: 50955
diff changeset
  6782
        apply (subst add_commute, rule dist_triangle)
3e5b67f85bf9 generalized theorem edelstein_fix to class metric_space
huffman
parents: 50955
diff changeset
  6783
        done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6784
      moreover
50970
3e5b67f85bf9 generalized theorem edelstein_fix to class metric_space
huffman
parents: 50955
diff changeset
  6785
      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
  6786
        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
  6787
        using seq_suble[OF r, of "Na+Nb+n"]
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6788
        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
  6789
      ultimately have False by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6790
    }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6791
    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
  6792
  note ab_fn = this
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6793
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6794
  have [simp]:"a = b" proof(rule ccontr)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6795
    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
  6796
    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
  6797
    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
  6798
      using lima limb unfolding LIMSEQ_def
93943da0a010 remove redundant lemma Lim_sequentially in favor of lemma LIMSEQ_def
huffman
parents: 44905
diff changeset
  6799
      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
  6800
    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
  6801
    have "dist (f (Suc n) x) (g a) \<le> dist (f n x) a"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6802
      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
  6803
    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
  6804
      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
  6805
    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
  6806
    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
  6807
      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
  6808
      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
  6809
      by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6810
  qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6811
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6812
  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
  6813
  { fix x y assume "x\<in>s" "y\<in>s" moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6814
    fix e::real assume "e>0" ultimately
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 44668
diff changeset
  6815
    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
  6816
  hence "continuous_on s g" unfolding continuous_on_iff by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6817
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6818
  hence "((snd \<circ> h \<circ> r) ---> g a) sequentially" unfolding continuous_on_sequentially
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6819
    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
  6820
    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
  6821
  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
  6822
    unfolding `a=b` and o_assoc by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6823
  moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6824
  { fix x assume "x\<in>s" "g x = x" "x\<noteq>a"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6825
    hence "False" using dist[THEN bspec[where x=a], THEN bspec[where x=x]]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6826
      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
  6827
  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
  6828
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6829
44131
5fc334b94e00 declare tendsto_const [intro] (accidentally removed in 230a8665c919)
huffman
parents: 44129
diff changeset
  6830
declare tendsto_const [intro] (* FIXME: move *)
5fc334b94e00 declare tendsto_const [intro] (accidentally removed in 230a8665c919)
huffman
parents: 44129
diff changeset
  6831
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6832
end