src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy
author hoelzl
Thu, 17 Jan 2013 12:26:54 +0100
changeset 50940 a7c273a83d27
parent 50939 ae7cd20ed118
child 50941 3690724028b1
permissions -rw-r--r--
group compactness-eq-seq-compactness lemmas together
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
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
    20
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
    21
  "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
    22
  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
    23
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    24
subsection {* Topological Basis *}
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    25
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    26
context topological_space
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    27
begin
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    28
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    29
definition "topological_basis B =
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    30
  ((\<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
    31
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    32
lemma topological_basis_iff:
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    33
  assumes "\<And>B'. B' \<in> B \<Longrightarrow> open B'"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    34
  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
    35
    (is "_ \<longleftrightarrow> ?rhs")
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    36
proof safe
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    37
  fix O' and x::'a
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    38
  assume H: "topological_basis B" "open O'" "x \<in> O'"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    39
  hence "(\<exists>B'\<subseteq>B. \<Union>B' = O')" by (simp add: topological_basis_def)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    40
  then obtain B' where "B' \<subseteq> B" "O' = \<Union>B'" by auto
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    41
  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
    42
next
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    43
  assume H: ?rhs
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    44
  show "topological_basis B" using assms unfolding topological_basis_def
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    45
  proof safe
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    46
    fix O'::"'a set" assume "open O'"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    47
    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
    48
      by (force intro: bchoice simp: Bex_def)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    49
    thus "\<exists>B'\<subseteq>B. \<Union>B' = O'"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    50
      by (auto intro: exI[where x="{f x |x. x \<in> O'}"])
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    51
  qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    52
qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    53
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    54
lemma topological_basisI:
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    55
  assumes "\<And>B'. B' \<in> B \<Longrightarrow> open B'"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    56
  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
    57
  shows "topological_basis B"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    58
  using assms by (subst topological_basis_iff) auto
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    59
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    60
lemma topological_basisE:
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    61
  fixes O'
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    62
  assumes "topological_basis B"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    63
  assumes "open O'"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    64
  assumes "x \<in> O'"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    65
  obtains B' where "B' \<in> B" "x \<in> B'" "B' \<subseteq> O'"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    66
proof atomize_elim
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    67
  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
    68
  with topological_basis_iff assms
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    69
  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
    70
qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    71
50094
84ddcf5364b4 allow arbitrary enumerations of basis in locale for generation of borel sets
immler
parents: 50087
diff changeset
    72
lemma topological_basis_open:
84ddcf5364b4 allow arbitrary enumerations of basis in locale for generation of borel sets
immler
parents: 50087
diff changeset
    73
  assumes "topological_basis B"
84ddcf5364b4 allow arbitrary enumerations of basis in locale for generation of borel sets
immler
parents: 50087
diff changeset
    74
  assumes "X \<in> B"
84ddcf5364b4 allow arbitrary enumerations of basis in locale for generation of borel sets
immler
parents: 50087
diff changeset
    75
  shows "open X"
84ddcf5364b4 allow arbitrary enumerations of basis in locale for generation of borel sets
immler
parents: 50087
diff changeset
    76
  using assms
84ddcf5364b4 allow arbitrary enumerations of basis in locale for generation of borel sets
immler
parents: 50087
diff changeset
    77
  by (simp add: topological_basis_def)
84ddcf5364b4 allow arbitrary enumerations of basis in locale for generation of borel sets
immler
parents: 50087
diff changeset
    78
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
    79
lemma basis_dense:
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
    80
  fixes B::"'a set set" and f::"'a set \<Rightarrow> 'a"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
    81
  assumes "topological_basis B"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
    82
  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
    83
  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
    84
proof (intro allI impI)
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
    85
  fix X::"'a set" assume "open X" "X \<noteq> {}"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
    86
  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
    87
  guess B' . note B' = this
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
    88
  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
    89
qed
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
    90
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    91
end
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    92
50882
a382bf90867e move prod instantiation of second_countable_topology to its definition
hoelzl
parents: 50881
diff changeset
    93
lemma topological_basis_prod:
a382bf90867e move prod instantiation of second_countable_topology to its definition
hoelzl
parents: 50881
diff changeset
    94
  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
    95
  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
    96
  unfolding topological_basis_def
a382bf90867e move prod instantiation of second_countable_topology to its definition
hoelzl
parents: 50881
diff changeset
    97
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
    98
  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
    99
  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
   100
  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
   101
    fix x y assume "(x, y) \<in> S"
a382bf90867e move prod instantiation of second_countable_topology to its definition
hoelzl
parents: 50881
diff changeset
   102
    from open_prod_elim[OF `open S` this]
a382bf90867e move prod instantiation of second_countable_topology to its definition
hoelzl
parents: 50881
diff changeset
   103
    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
   104
      by (metis mem_Sigma_iff)
a382bf90867e move prod instantiation of second_countable_topology to its definition
hoelzl
parents: 50881
diff changeset
   105
    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
   106
    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
   107
    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
   108
      by (intro UN_I[of "(A0, B0)"]) auto
a382bf90867e move prod instantiation of second_countable_topology to its definition
hoelzl
parents: 50881
diff changeset
   109
  qed auto
a382bf90867e move prod instantiation of second_countable_topology to its definition
hoelzl
parents: 50881
diff changeset
   110
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
   111
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   112
subsection {* Countable Basis *}
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   113
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   114
locale countable_basis =
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   115
  fixes B::"'a::topological_space set set"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   116
  assumes is_basis: "topological_basis B"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   117
  assumes countable_basis: "countable B"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   118
begin
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   119
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   120
lemma open_countable_basis_ex:
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   121
  assumes "open X"
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   122
  shows "\<exists>B' \<subseteq> B. X = Union B'"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   123
  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
   124
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   125
lemma open_countable_basisE:
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   126
  assumes "open X"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   127
  obtains B' where "B' \<subseteq> B" "X = Union B'"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   128
  using assms open_countable_basis_ex by (atomize_elim) simp
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   129
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   130
lemma countable_dense_exists:
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   131
  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
   132
proof -
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   133
  let ?f = "(\<lambda>B'. SOME x. x \<in> B')"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   134
  have "countable (?f ` B)" using countable_basis by simp
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   135
  with basis_dense[OF is_basis, of ?f] show ?thesis
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   136
    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
   137
qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   138
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   139
lemma countable_dense_setE:
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   140
  obtains D :: "'a set"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   141
  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
   142
  using countable_dense_exists by blast
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   143
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   144
text {* Construction of an increasing sequence approximating open sets,
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   145
  therefore basis which is closed under union. *}
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   146
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   147
definition union_closed_basis::"'a set set" where
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   148
  "union_closed_basis = (\<lambda>l. \<Union>set l) ` lists B"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   149
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   150
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
   151
proof (rule topological_basisI)
84ddcf5364b4 allow arbitrary enumerations of basis in locale for generation of borel sets
immler
parents: 50087
diff changeset
   152
  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
   153
  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
   154
  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
   155
    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
   156
next
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   157
  fix B' assume "B' \<in> union_closed_basis"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   158
  thus "open B'"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   159
    using topological_basis_open[OF is_basis]
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   160
    by (auto simp: union_closed_basis_def)
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   161
qed
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   162
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   163
lemma countable_union_closed_basis: "countable union_closed_basis"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   164
  unfolding union_closed_basis_def using countable_basis by simp
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   165
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   166
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
   167
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   168
lemma union_closed_basis_ex:
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   169
 assumes X: "X \<in> union_closed_basis"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   170
 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
   171
proof -
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   172
  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
   173
  thus ?thesis by auto
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   174
qed
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
lemma union_closed_basisE:
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   177
  assumes "X \<in> union_closed_basis"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   178
  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
   179
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   180
lemma union_closed_basisI:
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   181
  assumes "finite B'" "X = \<Union>B'" "B' \<subseteq> B"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   182
  shows "X \<in> union_closed_basis"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   183
proof -
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   184
  from finite_list[OF `finite B'`] guess l ..
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   185
  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
   186
qed
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   187
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   188
lemma empty_basisI[intro]: "{} \<in> union_closed_basis"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   189
  by (rule union_closed_basisI[of "{}"]) auto
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   190
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   191
lemma union_basisI[intro]:
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   192
  assumes "X \<in> union_closed_basis" "Y \<in> union_closed_basis"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   193
  shows "X \<union> Y \<in> union_closed_basis"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   194
  using assms by (auto intro: union_closed_basisI elim!:union_closed_basisE)
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   195
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   196
lemma open_imp_Union_of_incseq:
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   197
  assumes "open X"
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   198
  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
   199
proof -
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   200
  from open_countable_basis_ex[OF `open X`]
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   201
  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
   202
  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
   203
  show ?thesis
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   204
  proof cases
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   205
    assume "B' \<noteq> {}"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   206
    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
   207
    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
   208
    have "incseq S" by (force simp: S_def incseq_Suc_iff)
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   209
    moreover
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   210
    have "(\<Union>j. S j) = X" unfolding B'
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   211
    proof safe
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   212
      fix x X assume "X \<in> B'" "x \<in> X"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   213
      then obtain n where "X = from_nat_into B' n"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   214
        by (metis `countable B'` from_nat_into_surj)
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   215
      also have "\<dots> \<subseteq> S n" by (auto simp: S_def)
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   216
      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
   217
    next
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   218
      fix x n
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   219
      assume "x \<in> S n"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   220
      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
   221
        by (simp add: S_def)
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   222
      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
   223
      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
   224
      finally show "x \<in> \<Union>B'" .
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   225
    qed
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   226
    moreover have "range S \<subseteq> union_closed_basis" using B'
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   227
      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
   228
    ultimately show ?thesis by auto
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   229
  qed (auto simp: B')
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   230
qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   231
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   232
lemma open_incseqE:
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   233
  assumes "open X"
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   234
  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
   235
  using open_imp_Union_of_incseq assms by atomize_elim
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   236
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   237
end
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   238
50883
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   239
class first_countable_topology = topological_space +
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   240
  assumes first_countable_basis:
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   241
    "\<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
   242
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   243
lemma (in first_countable_topology) countable_basis_at_decseq:
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   244
  obtains A :: "nat \<Rightarrow> 'a set" where
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   245
    "\<And>i. open (A i)" "\<And>i. x \<in> (A i)"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   246
    "\<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
   247
proof atomize_elim
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   248
  from first_countable_basis[of x] obtain A
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   249
    where "countable A"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   250
    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
   251
    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
   252
  then have "A \<noteq> {}" by auto
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   253
  with `countable A` have r: "A = range (from_nat_into A)" by auto
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   254
  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
   255
  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
   256
      (\<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
   257
  proof (safe intro!: exI[of _ F])
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   258
    fix i
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   259
    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
   260
    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
   261
  next
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   262
    fix S assume "open S" "x \<in> S"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   263
    from incl[OF this] obtain i where "F i \<subseteq> S"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   264
      by (subst (asm) r) (auto simp: F_def)
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   265
    moreover have "\<And>j. i \<le> j \<Longrightarrow> F j \<subseteq> F i"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   266
      by (auto simp: F_def)
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   267
    ultimately show "eventually (\<lambda>i. F i \<subseteq> S) sequentially"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   268
      by (auto simp: eventually_sequentially)
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   269
  qed
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   270
qed
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   271
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   272
lemma (in first_countable_topology) first_countable_basisE:
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   273
  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
   274
    "\<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
   275
  using first_countable_basis[of x]
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   276
  by atomize_elim auto
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   277
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   278
instance prod :: (first_countable_topology, first_countable_topology) first_countable_topology
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   279
proof
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   280
  fix x :: "'a \<times> 'b"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   281
  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
   282
  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
   283
  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
   284
  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
   285
    fix a b assume x: "a \<in> A" "b \<in> B"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   286
    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
   287
      unfolding mem_Times_iff by (auto intro: open_Times)
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   288
  next
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   289
    fix S assume "open S" "x \<in> S"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   290
    from open_prod_elim[OF this] guess a' b' .
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   291
    moreover with A(4)[of a'] B(4)[of b']
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   292
    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
   293
    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
   294
      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
   295
  qed (simp add: A B)
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   296
qed
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   297
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   298
instance metric_space \<subseteq> first_countable_topology
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   299
proof
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   300
  fix x :: 'a
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   301
  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
   302
  proof (intro exI, safe)
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   303
    fix S assume "open S" "x \<in> S"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   304
    then obtain r where "0 < r" "{y. dist x y < r} \<subseteq> S"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   305
      by (auto simp: open_dist dist_commute subset_eq)
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   306
    moreover from reals_Archimedean[OF `0 < r`] guess n ..
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   307
    moreover
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   308
    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
   309
      by (auto simp: inverse_eq_divide)
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   310
    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
   311
      by auto
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   312
  qed (auto intro: open_ball)
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   313
qed
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   314
50881
ae630bab13da renamed countable_basis_space to second_countable_topology
hoelzl
parents: 50526
diff changeset
   315
class second_countable_topology = topological_space +
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   316
  assumes ex_countable_basis:
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   317
    "\<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
   318
50881
ae630bab13da renamed countable_basis_space to second_countable_topology
hoelzl
parents: 50526
diff changeset
   319
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
   320
  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
   321
50882
a382bf90867e move prod instantiation of second_countable_topology to its definition
hoelzl
parents: 50881
diff changeset
   322
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
   323
proof
a382bf90867e move prod instantiation of second_countable_topology to its definition
hoelzl
parents: 50881
diff changeset
   324
  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
   325
    using ex_countable_basis by auto
a382bf90867e move prod instantiation of second_countable_topology to its definition
hoelzl
parents: 50881
diff changeset
   326
  moreover
a382bf90867e move prod instantiation of second_countable_topology to its definition
hoelzl
parents: 50881
diff changeset
   327
  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
   328
    using ex_countable_basis by auto
a382bf90867e move prod instantiation of second_countable_topology to its definition
hoelzl
parents: 50881
diff changeset
   329
  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
   330
    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
   331
qed
a382bf90867e move prod instantiation of second_countable_topology to its definition
hoelzl
parents: 50881
diff changeset
   332
50883
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   333
instance second_countable_topology \<subseteq> first_countable_topology
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   334
proof
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   335
  fix x :: 'a
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   336
  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
   337
  then have B: "countable B" "topological_basis B"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   338
    using countable_basis is_basis
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   339
    by (auto simp: countable_basis is_basis)
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   340
  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
   341
    by (intro exI[of _ "{b\<in>B. x \<in> b}"])
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   342
       (fastforce simp: topological_space_class.topological_basis_def)
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   343
qed
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
   344
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   345
subsection {* Polish spaces *}
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   346
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   347
text {* Textbooks define Polish spaces as completely metrizable.
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   348
  We assume the topology to be complete for a given metric. *}
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   349
50881
ae630bab13da renamed countable_basis_space to second_countable_topology
hoelzl
parents: 50526
diff changeset
   350
class polish_space = complete_space + second_countable_topology
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   351
44517
68e8eb0ce8aa minimize imports
huffman
parents: 44516
diff changeset
   352
subsection {* General notion of a topology as a value *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   353
44170
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   354
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
   355
typedef 'a topology = "{L::('a set) \<Rightarrow> bool. istopology L}"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   356
  morphisms "openin" "topology"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   357
  unfolding istopology_def by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   358
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   359
lemma istopology_open_in[intro]: "istopology(openin U)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   360
  using openin[of U] by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   361
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   362
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
   363
  using topology_inverse[unfolded mem_Collect_eq] .
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   364
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   365
lemma topology_inverse_iff: "istopology U \<longleftrightarrow> openin (topology U) = U"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   366
  using topology_inverse[of U] istopology_open_in[of "topology U"] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   367
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   368
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
   369
proof-
49711
e5aaae7eadc9 tuned proofs;
wenzelm
parents: 48125
diff changeset
   370
  { assume "T1=T2"
e5aaae7eadc9 tuned proofs;
wenzelm
parents: 48125
diff changeset
   371
    hence "\<forall>S. openin T1 S \<longleftrightarrow> openin T2 S" by simp }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   372
  moreover
49711
e5aaae7eadc9 tuned proofs;
wenzelm
parents: 48125
diff changeset
   373
  { 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
   374
    hence "openin T1 = openin T2" by (simp add: fun_eq_iff)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   375
    hence "topology (openin T1) = topology (openin T2)" by simp
49711
e5aaae7eadc9 tuned proofs;
wenzelm
parents: 48125
diff changeset
   376
    hence "T1 = T2" unfolding openin_inverse .
e5aaae7eadc9 tuned proofs;
wenzelm
parents: 48125
diff changeset
   377
  }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   378
  ultimately show ?thesis by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   379
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   380
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   381
text{* Infer the "universe" from union of all sets in the topology. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   382
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   383
definition "topspace T =  \<Union>{S. openin T S}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   384
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   385
subsubsection {* Main properties of open sets *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   386
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   387
lemma openin_clauses:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   388
  fixes U :: "'a topology"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   389
  shows "openin U {}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   390
  "\<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
   391
  "\<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
   392
  using openin[of U] unfolding istopology_def mem_Collect_eq
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   393
  by fast+
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   394
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   395
lemma openin_subset[intro]: "openin U S \<Longrightarrow> S \<subseteq> topspace U"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   396
  unfolding topspace_def by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   397
lemma openin_empty[simp]: "openin U {}" by (simp add: openin_clauses)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   398
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   399
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
   400
  using openin_clauses by simp
06475a1547cb fix lots of looping simp calls and other warnings
huffman
parents: 36360
diff changeset
   401
06475a1547cb fix lots of looping simp calls and other warnings
huffman
parents: 36360
diff changeset
   402
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
   403
  using openin_clauses by simp
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_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
   406
  using openin_Union[of "{S,T}" U] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   407
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   408
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
   409
49711
e5aaae7eadc9 tuned proofs;
wenzelm
parents: 48125
diff changeset
   410
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
   411
  (is "?lhs \<longleftrightarrow> ?rhs")
36584
1535841fc2e9 prove lemma openin_subopen without using choice
huffman
parents: 36442
diff changeset
   412
proof
49711
e5aaae7eadc9 tuned proofs;
wenzelm
parents: 48125
diff changeset
   413
  assume ?lhs
e5aaae7eadc9 tuned proofs;
wenzelm
parents: 48125
diff changeset
   414
  then show ?rhs by auto
36584
1535841fc2e9 prove lemma openin_subopen without using choice
huffman
parents: 36442
diff changeset
   415
next
1535841fc2e9 prove lemma openin_subopen without using choice
huffman
parents: 36442
diff changeset
   416
  assume H: ?rhs
1535841fc2e9 prove lemma openin_subopen without using choice
huffman
parents: 36442
diff changeset
   417
  let ?t = "\<Union>{T. openin U T \<and> T \<subseteq> S}"
1535841fc2e9 prove lemma openin_subopen without using choice
huffman
parents: 36442
diff changeset
   418
  have "openin U ?t" by (simp add: openin_Union)
1535841fc2e9 prove lemma openin_subopen without using choice
huffman
parents: 36442
diff changeset
   419
  also have "?t = S" using H by auto
1535841fc2e9 prove lemma openin_subopen without using choice
huffman
parents: 36442
diff changeset
   420
  finally show "openin U S" .
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   421
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   422
49711
e5aaae7eadc9 tuned proofs;
wenzelm
parents: 48125
diff changeset
   423
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   424
subsubsection {* Closed sets *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   425
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   426
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
   427
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   428
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
   429
lemma closedin_empty[simp]: "closedin U {}" by (simp add: closedin_def)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   430
lemma closedin_topspace[intro,simp]:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   431
  "closedin U (topspace U)" by (simp add: closedin_def)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   432
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
   433
  by (auto simp add: Diff_Un closedin_def)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   434
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   435
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
   436
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
   437
  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
   438
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   439
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
   440
  using closedin_Inter[of "{S,T}" U] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   441
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   442
lemma Diff_Diff_Int: "A - (A - B) = A \<inter> B" by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   443
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
   444
  apply (auto simp add: closedin_def Diff_Diff_Int inf_absorb2)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   445
  apply (metis openin_subset subset_eq)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   446
  done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   447
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   448
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
   449
  by (simp add: openin_closedin_eq)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   450
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   451
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
   452
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   453
  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
   454
    by (auto simp add: topspace_def openin_subset)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   455
  then show ?thesis using oS cT by (auto simp add: closedin_def)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   456
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   457
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   458
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
   459
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   460
  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
   461
    by (auto simp add: topspace_def )
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   462
  then show ?thesis using oS cT by (auto simp add: openin_closedin_eq)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   463
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   464
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   465
subsubsection {* Subspace topology *}
44170
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   466
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   467
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
   468
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   469
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
   470
  (is "istopology ?L")
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   471
proof-
44170
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   472
  have "?L {}" by blast
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   473
  {fix A B assume A: "?L A" and B: "?L B"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   474
    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
   475
    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
   476
    then have "?L (A \<inter> B)" by blast}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   477
  moreover
44170
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   478
  {fix K assume K: "K \<subseteq> Collect ?L"
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   479
    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
   480
      apply (rule set_eqI)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   481
      apply (simp add: Ball_def image_iff)
44170
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   482
      by metis
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   483
    from K[unfolded th0 subset_image_iff]
44170
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   484
    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
   485
    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
   486
    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
   487
    ultimately have "?L (\<Union>K)" by blast}
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   488
  ultimately show ?thesis
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   489
    unfolding subset_eq mem_Collect_eq istopology_def by blast
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   490
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   491
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   492
lemma openin_subtopology:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   493
  "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
   494
  unfolding subtopology_def topology_inverse'[OF istopology_subtopology]
44170
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   495
  by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   496
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   497
lemma topspace_subtopology: "topspace(subtopology U V) = topspace U \<inter> V"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   498
  by (auto simp add: topspace_def openin_subtopology)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   499
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   500
lemma closedin_subtopology:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   501
  "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
   502
  unfolding closedin_def topspace_subtopology
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   503
  apply (simp add: openin_subtopology)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   504
  apply (rule iffI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   505
  apply clarify
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   506
  apply (rule_tac x="topspace U - T" in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   507
  by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   508
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   509
lemma openin_subtopology_refl: "openin (subtopology U V) V \<longleftrightarrow> V \<subseteq> topspace U"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   510
  unfolding openin_subtopology
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   511
  apply (rule iffI, clarify)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   512
  apply (frule openin_subset[of U])  apply blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   513
  apply (rule exI[where x="topspace U"])
49711
e5aaae7eadc9 tuned proofs;
wenzelm
parents: 48125
diff changeset
   514
  apply auto
e5aaae7eadc9 tuned proofs;
wenzelm
parents: 48125
diff changeset
   515
  done
e5aaae7eadc9 tuned proofs;
wenzelm
parents: 48125
diff changeset
   516
e5aaae7eadc9 tuned proofs;
wenzelm
parents: 48125
diff changeset
   517
lemma subtopology_superset:
e5aaae7eadc9 tuned proofs;
wenzelm
parents: 48125
diff changeset
   518
  assumes UV: "topspace U \<subseteq> V"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   519
  shows "subtopology U V = U"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   520
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   521
  {fix S
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   522
    {fix T assume T: "openin U T" "S = T \<inter> V"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   523
      from T openin_subset[OF T(1)] UV have eq: "S = T" by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   524
      have "openin U S" unfolding eq using T by blast}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   525
    moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   526
    {assume S: "openin U S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   527
      hence "\<exists>T. openin U T \<and> S = T \<inter> V"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   528
        using openin_subset[OF S] UV by auto}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   529
    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
   530
  then show ?thesis unfolding topology_eq openin_subtopology by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   531
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   532
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   533
lemma subtopology_topspace[simp]: "subtopology U (topspace U) = U"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   534
  by (simp add: subtopology_superset)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   535
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   536
lemma subtopology_UNIV[simp]: "subtopology U UNIV = U"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   537
  by (simp add: subtopology_superset)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   538
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   539
subsubsection {* The standard Euclidean topology *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   540
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   541
definition
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   542
  euclidean :: "'a::topological_space topology" where
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   543
  "euclidean = topology open"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   544
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   545
lemma open_openin: "open S \<longleftrightarrow> openin euclidean S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   546
  unfolding euclidean_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   547
  apply (rule cong[where x=S and y=S])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   548
  apply (rule topology_inverse[symmetric])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   549
  apply (auto simp add: istopology_def)
44170
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   550
  done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   551
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   552
lemma topspace_euclidean: "topspace euclidean = UNIV"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   553
  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
   554
  apply (rule set_eqI)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   555
  by (auto simp add: open_openin[symmetric])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   556
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   557
lemma topspace_euclidean_subtopology[simp]: "topspace (subtopology euclidean S) = S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   558
  by (simp add: topspace_euclidean topspace_subtopology)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   559
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   560
lemma closed_closedin: "closed S \<longleftrightarrow> closedin euclidean S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   561
  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
   562
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   563
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
   564
  by (simp add: open_openin openin_subopen[symmetric])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   565
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   566
text {* Basic "localization" results are handy for connectedness. *}
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   567
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   568
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
   569
  by (auto simp add: openin_subtopology open_openin[symmetric])
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   570
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   571
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
   572
  by (auto simp add: openin_open)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   573
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   574
lemma open_openin_trans[trans]:
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   575
 "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
   576
  by (metis Int_absorb1  openin_open_Int)
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 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
   579
  by (auto simp add: openin_open)
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 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
   582
  by (simp add: closedin_subtopology closed_closedin Int_ac)
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 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
   585
  by (metis closedin_closed)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   586
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   587
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
   588
  apply (subgoal_tac "S \<inter> T = T" )
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   589
  apply auto
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   590
  apply (frule closedin_closed_Int[of T S])
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   591
  by simp
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   592
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   593
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
   594
  by (auto simp add: closedin_closed)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   595
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   596
lemma openin_euclidean_subtopology_iff:
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   597
  fixes S U :: "'a::metric_space set"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   598
  shows "openin (subtopology euclidean U) S
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   599
  \<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
   600
proof
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   601
  assume ?lhs thus ?rhs unfolding openin_open open_dist by blast
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   602
next
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   603
  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
   604
  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
   605
    unfolding T_def
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   606
    apply clarsimp
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   607
    apply (rule_tac x="d - dist x a" in exI)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   608
    apply (clarsimp simp add: less_diff_eq)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   609
    apply (erule rev_bexI)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   610
    apply (rule_tac x=d in exI, clarify)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   611
    apply (erule le_less_trans [OF dist_triangle])
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   612
    done
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   613
  assume ?rhs hence 2: "S = U \<inter> T"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   614
    unfolding T_def
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   615
    apply auto
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   616
    apply (drule (1) bspec, erule rev_bexI)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   617
    apply auto
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   618
    done
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   619
  from 1 2 show ?lhs
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   620
    unfolding openin_open open_dist by fast
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   621
qed
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   622
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   623
text {* These "transitivity" results are handy too *}
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   624
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   625
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
   626
  \<Longrightarrow> openin (subtopology euclidean U) S"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   627
  unfolding open_openin openin_open by blast
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   628
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   629
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
   630
  by (auto simp add: openin_open intro: openin_trans)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   631
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   632
lemma closedin_trans[trans]:
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   633
 "closedin (subtopology euclidean T) S \<Longrightarrow>
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   634
           closedin (subtopology euclidean U) T
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   635
           ==> closedin (subtopology euclidean U) S"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   636
  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
   637
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   638
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
   639
  by (auto simp add: closedin_closed intro: closedin_trans)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   640
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   641
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   642
subsection {* Open and closed balls *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   643
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   644
definition
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   645
  ball :: "'a::metric_space \<Rightarrow> real \<Rightarrow> 'a set" where
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   646
  "ball x e = {y. dist x y < e}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   647
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   648
definition
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   649
  cball :: "'a::metric_space \<Rightarrow> real \<Rightarrow> 'a set" where
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   650
  "cball x e = {y. dist x y \<le> e}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   651
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
   652
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
   653
  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
   654
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
   655
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
   656
  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
   657
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
   658
lemma mem_ball_0:
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   659
  fixes x :: "'a::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   660
  shows "x \<in> ball 0 e \<longleftrightarrow> norm x < e"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   661
  by (simp add: dist_norm)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   662
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
   663
lemma mem_cball_0:
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   664
  fixes x :: "'a::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   665
  shows "x \<in> cball 0 e \<longleftrightarrow> norm x \<le> e"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   666
  by (simp add: dist_norm)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   667
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
   668
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
   669
  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
   670
714100f5fda4 remove mem_(c)ball_0 and centre_in_(c)ball from simpset, as rules mem_(c)ball always match instead
huffman
parents: 45548
diff changeset
   671
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
   672
  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
   673
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   674
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
   675
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
   676
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
   677
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
   678
  by (simp add: set_eq_iff) arith
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   679
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   680
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
   681
  by (simp add: set_eq_iff)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   682
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   683
lemma diff_less_iff: "(a::real) - b > 0 \<longleftrightarrow> a > b"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   684
  "(a::real) - b < 0 \<longleftrightarrow> a < b"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   685
  "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
   686
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
   687
  "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
   688
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   689
lemma open_ball[intro, simp]: "open (ball x e)"
44170
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   690
  unfolding open_dist ball_def mem_Collect_eq Ball_def
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   691
  unfolding dist_commute
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   692
  apply clarify
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   693
  apply (rule_tac x="e - dist xa x" in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   694
  using dist_triangle_alt[where z=x]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   695
  apply (clarsimp simp add: diff_less_iff)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   696
  apply atomize
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   697
  apply (erule_tac x="y" in allE)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   698
  apply (erule_tac x="xa" in allE)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   699
  by arith
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   700
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   701
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
   702
  unfolding open_dist subset_eq mem_ball Ball_def dist_commute ..
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   703
33714
eb2574ac4173 Added new lemmas to Euclidean Space by Robert Himmelmann
hoelzl
parents: 33324
diff changeset
   704
lemma openE[elim?]:
eb2574ac4173 Added new lemmas to Euclidean Space by Robert Himmelmann
hoelzl
parents: 33324
diff changeset
   705
  assumes "open S" "x\<in>S" 
eb2574ac4173 Added new lemmas to Euclidean Space by Robert Himmelmann
hoelzl
parents: 33324
diff changeset
   706
  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
   707
  using assms unfolding open_contains_ball by auto
eb2574ac4173 Added new lemmas to Euclidean Space by Robert Himmelmann
hoelzl
parents: 33324
diff changeset
   708
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   709
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
   710
  by (metis open_contains_ball subset_eq centre_in_ball)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   711
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   712
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
   713
  unfolding mem_ball set_eq_iff
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   714
  apply (simp add: not_less)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   715
  by (metis zero_le_dist order_trans dist_self)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   716
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   717
lemma ball_empty[intro]: "e \<le> 0 ==> ball x e = {}" by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   718
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
   719
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
   720
  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
   721
  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
   722
  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
   723
  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
   724
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
   725
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
   726
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   727
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
   728
  fixes x :: "'a\<Colon>euclidean_space"
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   729
  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
   730
  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
   731
proof -
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   732
  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
   733
  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
   734
  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
   735
  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
   736
    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
   737
  qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   738
  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
   739
  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
   740
  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
   741
    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
   742
  qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   743
  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
   744
  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
   745
  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
   746
  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
   747
    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
   748
    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
   749
      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
   750
    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
   751
    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
   752
      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
   753
      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
   754
      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
   755
      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
   756
      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
   757
      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
   758
        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
   759
      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
   760
        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
   761
      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
   762
        by (simp add: power_divide)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   763
    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
   764
    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
   765
    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
   766
  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
   767
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
   768
 
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
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
   770
  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
   771
  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
   772
  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
   773
  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
   774
  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
   775
  shows "M = (\<Union>f\<in>I. box (a' f) (b' f))"
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   776
proof safe
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   777
  fix x assume "x \<in> M"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   778
  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
   779
    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
   780
  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
   781
    "\<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
   782
    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
   783
  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
   784
     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
   785
        (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
   786
qed (auto simp: I_def)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   787
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   788
subsection{* Connectedness *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   789
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   790
definition "connected S \<longleftrightarrow>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   791
  ~(\<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
   792
  \<and> ~(e1 \<inter> S = {}) \<and> ~(e2 \<inter> S = {}))"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   793
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   794
lemma connected_local:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   795
 "connected S \<longleftrightarrow> ~(\<exists>e1 e2.
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   796
                 openin (subtopology euclidean S) e1 \<and>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   797
                 openin (subtopology euclidean S) e2 \<and>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   798
                 S \<subseteq> e1 \<union> e2 \<and>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   799
                 e1 \<inter> e2 = {} \<and>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   800
                 ~(e1 = {}) \<and>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   801
                 ~(e2 = {}))"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   802
unfolding connected_def openin_open by (safe, blast+)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   803
34105
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
   804
lemma exists_diff:
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
   805
  fixes P :: "'a set \<Rightarrow> bool"
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
   806
  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
   807
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   808
  {assume "?lhs" hence ?rhs by blast }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   809
  moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   810
  {fix S assume H: "P S"
34105
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
   811
    have "S = - (- S)" by auto
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
   812
    with H have "P (- (- S))" by metis }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   813
  ultimately show ?thesis by metis
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   814
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   815
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   816
lemma connected_clopen: "connected S \<longleftrightarrow>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   817
        (\<forall>T. openin (subtopology euclidean S) T \<and>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   818
            closedin (subtopology euclidean S) T \<longrightarrow> T = {} \<or> T = S)" (is "?lhs \<longleftrightarrow> ?rhs")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   819
proof-
34105
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
   820
  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
   821
    unfolding connected_def openin_open closedin_closed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   822
    apply (subst exists_diff) by blast
34105
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
   823
  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
   824
    (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
   825
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   826
  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
   827
    (is "_ \<longleftrightarrow> \<not> (\<exists>t' t. ?Q t' t)")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   828
    unfolding connected_def openin_open closedin_closed by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   829
  {fix e2
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   830
    {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
   831
        by auto}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   832
    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
   833
  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
   834
  then show ?thesis unfolding th0 th1 by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   835
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   836
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   837
lemma connected_empty[simp, intro]: "connected {}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   838
  by (simp add: connected_def)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   839
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   840
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   841
subsection{* Limit points *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   842
44207
ea99698c2070 locale-ize some definitions, so perfect_space and heine_borel can inherit from the proper superclasses
huffman
parents: 44170
diff changeset
   843
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
   844
  islimpt:: "'a \<Rightarrow> 'a set \<Rightarrow> bool" (infixr "islimpt" 60) where
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   845
  "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
   846
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   847
lemma islimptI:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   848
  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
   849
  shows "x islimpt S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   850
  using assms unfolding islimpt_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   851
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   852
lemma islimptE:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   853
  assumes "x islimpt S" and "x \<in> T" and "open T"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   854
  obtains y where "y \<in> S" and "y \<in> T" and "y \<noteq> x"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   855
  using assms unfolding islimpt_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   856
44584
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
   857
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
   858
  unfolding islimpt_def eventually_at_topological by auto
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
   859
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
   860
lemma islimpt_subset: "\<lbrakk>x islimpt S; S \<subseteq> T\<rbrakk> \<Longrightarrow> x islimpt T"
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
   861
  unfolding islimpt_def by fast
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   862
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   863
lemma islimpt_approachable:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   864
  fixes x :: "'a::metric_space"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   865
  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
   866
  unfolding islimpt_iff_eventually eventually_at by fast
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   867
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   868
lemma islimpt_approachable_le:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   869
  fixes x :: "'a::metric_space"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   870
  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
   871
  unfolding islimpt_approachable
44584
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
   872
  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
   873
    THEN arg_cong [where f=Not]]
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
   874
  by (simp add: Bex_def conj_commute conj_left_commute)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   875
44571
bd91b77c4cd6 move class perfect_space into RealVector.thy;
huffman
parents: 44568
diff changeset
   876
lemma islimpt_UNIV_iff: "x islimpt UNIV \<longleftrightarrow> \<not> open {x}"
bd91b77c4cd6 move class perfect_space into RealVector.thy;
huffman
parents: 44568
diff changeset
   877
  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
   878
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   879
text {* A perfect space has no isolated points. *}
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   880
44571
bd91b77c4cd6 move class perfect_space into RealVector.thy;
huffman
parents: 44568
diff changeset
   881
lemma islimpt_UNIV [simp, intro]: "(x::'a::perfect_space) islimpt UNIV"
bd91b77c4cd6 move class perfect_space into RealVector.thy;
huffman
parents: 44568
diff changeset
   882
  unfolding islimpt_UNIV_iff by (rule not_open_singleton)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   883
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   884
lemma perfect_choose_dist:
44072
5b970711fb39 class perfect_space inherits from topological_space;
huffman
parents: 43338
diff changeset
   885
  fixes x :: "'a::{perfect_space, metric_space}"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   886
  shows "0 < r \<Longrightarrow> \<exists>a. a \<noteq> x \<and> dist a x < r"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   887
using islimpt_UNIV [of x]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   888
by (simp add: islimpt_approachable)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   889
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   890
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
   891
  unfolding closed_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   892
  apply (subst open_subopen)
34105
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
   893
  apply (simp add: islimpt_def subset_eq)
44170
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   894
  by (metis ComplE ComplI)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   895
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   896
lemma islimpt_EMPTY[simp]: "\<not> x islimpt {}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   897
  unfolding islimpt_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   898
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   899
lemma finite_set_avoid:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   900
  fixes a :: "'a::metric_space"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   901
  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
   902
proof(induct rule: finite_induct[OF fS])
41863
e5104b436ea1 removed dependency on Dense_Linear_Order
boehmes
parents: 41413
diff changeset
   903
  case 1 thus ?case by (auto intro: zero_less_one)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   904
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   905
  case (2 x F)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   906
  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
   907
  {assume "x = a" hence ?case using d by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   908
  moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   909
  {assume xa: "x\<noteq>a"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   910
    let ?d = "min d (dist a x)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   911
    have dp: "?d > 0" using xa d(1) using dist_nz by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   912
    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
   913
    with dp xa have ?case by(auto intro!: exI[where x="?d"]) }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   914
  ultimately show ?case by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   915
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   916
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   917
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
   918
  by (simp add: islimpt_iff_eventually eventually_conj_iff)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   919
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   920
lemma discrete_imp_closed:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   921
  fixes S :: "'a::metric_space set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   922
  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
   923
  shows "closed S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   924
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   925
  {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
   926
    from e have e2: "e/2 > 0" by arith
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   927
    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
   928
    let ?m = "min (e/2) (dist x y) "
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   929
    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
   930
    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
   931
    have th: "dist z y < e" using z y
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   932
      by (intro dist_triangle_lt [where z=x], simp)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   933
    from d[rule_format, OF y(1) z(1) th] y z
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   934
    have False by (auto simp add: dist_commute)}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   935
  then show ?thesis by (metis islimpt_approachable closed_limpt [where 'a='a])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   936
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   937
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   938
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   939
subsection {* Interior of a Set *}
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   940
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   941
definition "interior S = \<Union>{T. open T \<and> T \<subseteq> S}"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   942
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   943
lemma interiorI [intro?]:
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   944
  assumes "open T" and "x \<in> T" and "T \<subseteq> S"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   945
  shows "x \<in> interior S"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   946
  using assms unfolding interior_def by fast
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   947
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   948
lemma interiorE [elim?]:
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   949
  assumes "x \<in> interior S"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   950
  obtains T where "open T" and "x \<in> T" and "T \<subseteq> S"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   951
  using assms unfolding interior_def by fast
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   952
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   953
lemma open_interior [simp, intro]: "open (interior S)"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   954
  by (simp add: interior_def open_Union)
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   955
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   956
lemma interior_subset: "interior S \<subseteq> S"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   957
  by (auto simp add: interior_def)
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   958
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   959
lemma interior_maximal: "T \<subseteq> S \<Longrightarrow> open T \<Longrightarrow> T \<subseteq> interior S"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   960
  by (auto simp add: interior_def)
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   961
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   962
lemma interior_open: "open S \<Longrightarrow> interior S = S"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   963
  by (intro equalityI interior_subset interior_maximal subset_refl)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   964
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   965
lemma interior_eq: "interior S = S \<longleftrightarrow> open S"
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   966
  by (metis open_interior interior_open)
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   967
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   968
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
   969
  by (metis interior_maximal interior_subset subset_trans)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   970
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   971
lemma interior_empty [simp]: "interior {} = {}"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   972
  using open_empty by (rule interior_open)
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   973
44522
2f7e9d890efe rename subset_{interior,closure} to {interior,closure}_mono;
huffman
parents: 44519
diff changeset
   974
lemma interior_UNIV [simp]: "interior UNIV = UNIV"
2f7e9d890efe rename subset_{interior,closure} to {interior,closure}_mono;
huffman
parents: 44519
diff changeset
   975
  using open_UNIV by (rule interior_open)
2f7e9d890efe rename subset_{interior,closure} to {interior,closure}_mono;
huffman
parents: 44519
diff changeset
   976
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   977
lemma interior_interior [simp]: "interior (interior S) = interior S"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   978
  using open_interior by (rule interior_open)
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   979
44522
2f7e9d890efe rename subset_{interior,closure} to {interior,closure}_mono;
huffman
parents: 44519
diff changeset
   980
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
   981
  by (auto simp add: interior_def)
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   982
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   983
lemma interior_unique:
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   984
  assumes "T \<subseteq> S" and "open T"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   985
  assumes "\<And>T'. T' \<subseteq> S \<Longrightarrow> open T' \<Longrightarrow> T' \<subseteq> T"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   986
  shows "interior S = T"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   987
  by (intro equalityI assms interior_subset open_interior interior_maximal)
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   988
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   989
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
   990
  by (intro equalityI Int_mono Int_greatest interior_mono Int_lower1
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   991
    Int_lower2 interior_maximal interior_subset open_Int open_interior)
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   992
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   993
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
   994
  using open_contains_ball_eq [where S="interior S"]
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   995
  by (simp add: open_subset_interior)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   996
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   997
lemma interior_limit_point [intro]:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   998
  fixes x :: "'a::perfect_space"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   999
  assumes x: "x \<in> interior S" shows "x islimpt S"
44072
5b970711fb39 class perfect_space inherits from topological_space;
huffman
parents: 43338
diff changeset
  1000
  using x islimpt_UNIV [of x]
5b970711fb39 class perfect_space inherits from topological_space;
huffman
parents: 43338
diff changeset
  1001
  unfolding interior_def islimpt_def
5b970711fb39 class perfect_space inherits from topological_space;
huffman
parents: 43338
diff changeset
  1002
  apply (clarsimp, rename_tac T T')
5b970711fb39 class perfect_space inherits from topological_space;
huffman
parents: 43338
diff changeset
  1003
  apply (drule_tac x="T \<inter> T'" in spec)
5b970711fb39 class perfect_space inherits from topological_space;
huffman
parents: 43338
diff changeset
  1004
  apply (auto simp add: open_Int)
5b970711fb39 class perfect_space inherits from topological_space;
huffman
parents: 43338
diff changeset
  1005
  done
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_closed_Un_empty_interior:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1008
  assumes cS: "closed S" and iT: "interior T = {}"
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1009
  shows "interior (S \<union> T) = interior S"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1010
proof
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1011
  show "interior S \<subseteq> interior (S \<union> T)"
44522
2f7e9d890efe rename subset_{interior,closure} to {interior,closure}_mono;
huffman
parents: 44519
diff changeset
  1012
    by (rule interior_mono, rule Un_upper1)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1013
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1014
  show "interior (S \<union> T) \<subseteq> interior S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1015
  proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1016
    fix x assume "x \<in> interior (S \<union> T)"
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1017
    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
  1018
    show "x \<in> interior S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1019
    proof (rule ccontr)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1020
      assume "x \<notin> interior S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1021
      with `x \<in> R` `open R` obtain y where "y \<in> R - S"
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1022
        unfolding interior_def by fast
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1023
      from `open R` `closed S` have "open (R - S)" by (rule open_Diff)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1024
      from `R \<subseteq> S \<union> T` have "R - S \<subseteq> T" by fast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1025
      from `y \<in> R - S` `open (R - S)` `R - S \<subseteq> T` `interior T = {}`
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1026
      show "False" unfolding interior_def by fast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1027
    qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1028
  qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1029
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1030
44365
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1031
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
  1032
proof (rule interior_unique)
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1033
  show "interior A \<times> interior B \<subseteq> A \<times> B"
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1034
    by (intro Sigma_mono interior_subset)
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1035
  show "open (interior A \<times> interior B)"
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1036
    by (intro open_Times open_interior)
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1037
  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
  1038
  proof (safe)
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1039
    fix x y assume "(x, y) \<in> T"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1040
    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
  1041
      using `open T` unfolding open_prod_def by fast
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1042
    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
  1043
      using `T \<subseteq> A \<times> B` by auto
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1044
    thus "x \<in> interior A" and "y \<in> interior B"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1045
      by (auto intro: interiorI)
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1046
  qed
44365
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1047
qed
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1048
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1049
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1050
subsection {* Closure of a Set *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1051
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1052
definition "closure S = S \<union> {x | x. x islimpt S}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1053
44518
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
  1054
lemma interior_closure: "interior S = - (closure (- S))"
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
  1055
  unfolding interior_def closure_def islimpt_def by auto
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
  1056
34105
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
  1057
lemma closure_interior: "closure S = - interior (- S)"
44518
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
  1058
  unfolding interior_closure by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1059
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1060
lemma closed_closure[simp, intro]: "closed (closure S)"
44518
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
  1061
  unfolding closure_interior by (simp add: closed_Compl)
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
  1062
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
  1063
lemma closure_subset: "S \<subseteq> closure S"
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
  1064
  unfolding closure_def by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1065
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1066
lemma closure_hull: "closure S = closed hull S"
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1067
  unfolding hull_def closure_interior interior_def by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1068
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1069
lemma closure_eq: "closure S = S \<longleftrightarrow> closed S"
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1070
  unfolding closure_hull using closed_Inter by (rule hull_eq)
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1071
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1072
lemma closure_closed [simp]: "closed S \<Longrightarrow> closure S = S"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1073
  unfolding closure_eq .
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1074
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1075
lemma closure_closure [simp]: "closure (closure S) = closure S"
44518
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
  1076
  unfolding closure_hull by (rule hull_hull)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1077
44522
2f7e9d890efe rename subset_{interior,closure} to {interior,closure}_mono;
huffman
parents: 44519
diff changeset
  1078
lemma closure_mono: "S \<subseteq> T \<Longrightarrow> closure S \<subseteq> closure T"
44518
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
  1079
  unfolding closure_hull by (rule hull_mono)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1080
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1081
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
  1082
  unfolding closure_hull by (rule hull_minimal)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1083
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1084
lemma closure_unique:
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1085
  assumes "S \<subseteq> T" and "closed T"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1086
  assumes "\<And>T'. S \<subseteq> T' \<Longrightarrow> closed T' \<Longrightarrow> T \<subseteq> T'"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1087
  shows "closure S = T"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1088
  using assms unfolding closure_hull by (rule hull_unique)
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1089
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1090
lemma closure_empty [simp]: "closure {} = {}"
44518
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
  1091
  using closed_empty by (rule closure_closed)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1092
44522
2f7e9d890efe rename subset_{interior,closure} to {interior,closure}_mono;
huffman
parents: 44519
diff changeset
  1093
lemma closure_UNIV [simp]: "closure UNIV = UNIV"
44518
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
  1094
  using closed_UNIV by (rule closure_closed)
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
  1095
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
  1096
lemma closure_union [simp]: "closure (S \<union> T) = closure S \<union> closure T"
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
  1097
  unfolding closure_interior by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1098
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1099
lemma closure_eq_empty: "closure S = {} \<longleftrightarrow> S = {}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1100
  using closure_empty closure_subset[of S]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1101
  by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1102
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1103
lemma closure_subset_eq: "closure S \<subseteq> S \<longleftrightarrow> closed S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1104
  using closure_eq[of S] closure_subset[of S]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1105
  by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1106
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1107
lemma open_inter_closure_eq_empty:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1108
  "open S \<Longrightarrow> (S \<inter> closure T) = {} \<longleftrightarrow> S \<inter> T = {}"
34105
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
  1109
  using open_subset_interior[of S "- T"]
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
  1110
  using interior_subset[of "- T"]
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1111
  unfolding closure_interior
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1112
  by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1113
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1114
lemma open_inter_closure_subset:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1115
  "open S \<Longrightarrow> (S \<inter> (closure T)) \<subseteq> closure(S \<inter> T)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1116
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1117
  fix x
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1118
  assume as: "open S" "x \<in> S \<inter> closure T"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1119
  { assume *:"x islimpt T"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1120
    have "x islimpt (S \<inter> T)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1121
    proof (rule islimptI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1122
      fix A
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1123
      assume "x \<in> A" "open A"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1124
      with as have "x \<in> A \<inter> S" "open (A \<inter> S)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1125
        by (simp_all add: open_Int)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1126
      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
  1127
        by (rule islimptE)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1128
      hence "y \<in> S \<inter> T" "y \<in> A \<and> y \<noteq> x"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1129
        by simp_all
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1130
      thus "\<exists>y\<in>(S \<inter> T). y \<in> A \<and> y \<noteq> x" ..
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1131
    qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1132
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1133
  then show "x \<in> closure (S \<inter> T)" using as
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1134
    unfolding closure_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1135
    by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1136
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1137
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1138
lemma closure_complement: "closure (- S) = - interior S"
44518
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
  1139
  unfolding closure_interior by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1140
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1141
lemma interior_complement: "interior (- S) = - closure S"
44518
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
  1142
  unfolding closure_interior by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1143
44365
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1144
lemma closure_Times: "closure (A \<times> B) = closure A \<times> closure B"
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1145
proof (rule closure_unique)
44365
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1146
  show "A \<times> B \<subseteq> closure A \<times> closure B"
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1147
    by (intro Sigma_mono closure_subset)
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1148
  show "closed (closure A \<times> closure B)"
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1149
    by (intro closed_Times closed_closure)
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1150
  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
  1151
    apply (simp add: closed_def open_prod_def, clarify)
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1152
    apply (rule ccontr)
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1153
    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
  1154
    apply (simp add: closure_interior interior_def)
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1155
    apply (drule_tac x=C in spec)
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1156
    apply (drule_tac x=D in spec)
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1157
    apply auto
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1158
    done
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1159
qed
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1160
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1161
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1162
subsection {* Frontier (aka boundary) *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1163
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1164
definition "frontier S = closure S - interior S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1165
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1166
lemma frontier_closed: "closed(frontier S)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1167
  by (simp add: frontier_def closed_Diff)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1168
34105
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
  1169
lemma frontier_closures: "frontier S = (closure S) \<inter> (closure(- S))"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1170
  by (auto simp add: frontier_def interior_closure)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1171
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1172
lemma frontier_straddle:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1173
  fixes a :: "'a::metric_space"
44909
1f5d6eb73549 shorten proof of frontier_straddle
huffman
parents: 44907
diff changeset
  1174
  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
  1175
  unfolding frontier_def closure_interior
1f5d6eb73549 shorten proof of frontier_straddle
huffman
parents: 44907
diff changeset
  1176
  by (auto simp add: mem_interior subset_eq ball_def)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1177
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1178
lemma frontier_subset_closed: "closed S \<Longrightarrow> frontier S \<subseteq> S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1179
  by (metis frontier_def closure_closed Diff_subset)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1180
34964
4e8be3c04d37 Replaced vec1 and dest_vec1 by abbreviation.
hoelzl
parents: 34291
diff changeset
  1181
lemma frontier_empty[simp]: "frontier {} = {}"
36362
06475a1547cb fix lots of looping simp calls and other warnings
huffman
parents: 36360
diff changeset
  1182
  by (simp add: frontier_def)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1183
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1184
lemma frontier_subset_eq: "frontier S \<subseteq> S \<longleftrightarrow> closed S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1185
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1186
  { assume "frontier S \<subseteq> S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1187
    hence "closure S \<subseteq> S" using interior_subset unfolding frontier_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1188
    hence "closed S" using closure_subset_eq by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1189
  }
36362
06475a1547cb fix lots of looping simp calls and other warnings
huffman
parents: 36360
diff changeset
  1190
  thus ?thesis using frontier_subset_closed[of S] ..
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1191
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1192
34105
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
  1193
lemma frontier_complement: "frontier(- S) = frontier S"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1194
  by (auto simp add: frontier_def closure_complement interior_complement)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1195
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1196
lemma frontier_disjoint_eq: "frontier S \<inter> S = {} \<longleftrightarrow> open S"
34105
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
  1197
  using frontier_complement frontier_subset_eq[of "- S"]
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
  1198
  unfolding open_closed by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1199
44081
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44076
diff changeset
  1200
subsection {* Filters and the ``eventually true'' quantifier *}
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44076
diff changeset
  1201
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1202
definition
44081
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44076
diff changeset
  1203
  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
  1204
    (infixr "indirection" 70) where
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1205
  "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
  1206
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  1207
text {* Identify Trivial limits, where we can't approach arbitrarily closely. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1208
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1209
lemma trivial_limit_within:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1210
  shows "trivial_limit (at a within S) \<longleftrightarrow> \<not> a islimpt S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1211
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1212
  assume "trivial_limit (at a within S)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1213
  thus "\<not> a islimpt S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1214
    unfolding trivial_limit_def
36358
246493d61204 define nets directly as filters, instead of as filter bases
huffman
parents: 36336
diff changeset
  1215
    unfolding eventually_within eventually_at_topological
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1216
    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
  1217
    apply (clarsimp simp add: set_eq_iff)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1218
    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
  1219
    apply (clarsimp, drule_tac x=y in bspec, simp_all)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1220
    done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1221
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1222
  assume "\<not> a islimpt S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1223
  thus "trivial_limit (at a within 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
36358
246493d61204 define nets directly as filters, instead of as filter bases
huffman
parents: 36336
diff changeset
  1227
    apply clarsimp
246493d61204 define nets directly as filters, instead of as filter bases
huffman
parents: 36336
diff changeset
  1228
    apply (rule_tac x=T in exI)
246493d61204 define nets directly as filters, instead of as filter bases
huffman
parents: 36336
diff changeset
  1229
    apply auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1230
    done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1231
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1232
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1233
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
  1234
  using trivial_limit_within [of a UNIV] by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1235
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1236
lemma trivial_limit_at:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1237
  fixes a :: "'a::perfect_space"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1238
  shows "\<not> trivial_limit (at a)"
44571
bd91b77c4cd6 move class perfect_space into RealVector.thy;
huffman
parents: 44568
diff changeset
  1239
  by (rule at_neq_bot)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1240
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1241
lemma trivial_limit_at_infinity:
44081
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44076
diff changeset
  1242
  "\<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
  1243
  unfolding trivial_limit_def eventually_at_infinity
246493d61204 define nets directly as filters, instead of as filter bases
huffman
parents: 36336
diff changeset
  1244
  apply clarsimp
44072
5b970711fb39 class perfect_space inherits from topological_space;
huffman
parents: 43338
diff changeset
  1245
  apply (subgoal_tac "\<exists>x::'a. x \<noteq> 0", clarify)
5b970711fb39 class perfect_space inherits from topological_space;
huffman
parents: 43338
diff changeset
  1246
   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
  1247
  apply (cut_tac islimpt_UNIV [of "0::'a", unfolded islimpt_def])
5b970711fb39 class perfect_space inherits from topological_space;
huffman
parents: 43338
diff changeset
  1248
  apply (drule_tac x=UNIV in spec, simp)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1249
  done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1250
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  1251
text {* Some property holds "sufficiently close" to the limit point. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1252
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1253
lemma eventually_at: (* FIXME: this replaces Limits.eventually_at *)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1254
  "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
  1255
unfolding eventually_at dist_nz by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1256
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
  1257
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
  1258
  "eventually P (at a within S) \<longleftrightarrow>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1259
        (\<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
  1260
  by (rule eventually_within_less)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1261
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1262
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
  1263
  unfolding trivial_limit_def
246493d61204 define nets directly as filters, instead of as filter bases
huffman
parents: 36336
diff changeset
  1264
  by (auto elim: eventually_rev_mp)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1265
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1266
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
  1267
  by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1268
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1269
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
  1270
  by (simp add: filter_eq_iff)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1271
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1272
text{* Combining theorems for "eventually" *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1273
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1274
lemma eventually_rev_mono:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1275
  "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
  1276
using eventually_mono [of P Q] by fast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1277
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1278
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
  1279
  by (simp add: eventually_False)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1280
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1281
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  1282
subsection {* Limits *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1283
44081
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44076
diff changeset
  1284
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
  1285
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44076
diff changeset
  1286
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
  1287
  where "Lim A f = (THE l. (f ---> l) A)"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1288
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1289
lemma Lim:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1290
 "(f ---> l) net \<longleftrightarrow>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1291
        trivial_limit net \<or>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1292
        (\<forall>e>0. eventually (\<lambda>x. dist (f x) l < e) net)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1293
  unfolding tendsto_iff trivial_limit_eq by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1294
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1295
text{* Show that they yield usual definitions in the various cases. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1296
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1297
lemma Lim_within_le: "(f ---> l)(at a within S) \<longleftrightarrow>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1298
           (\<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
  1299
  by (auto simp add: tendsto_iff eventually_within_le)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1300
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1301
lemma Lim_within: "(f ---> l) (at a within S) \<longleftrightarrow>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1302
        (\<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
  1303
  by (auto simp add: tendsto_iff eventually_within)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1304
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1305
lemma Lim_at: "(f ---> l) (at a) \<longleftrightarrow>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1306
        (\<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
  1307
  by (auto simp add: tendsto_iff eventually_at)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1308
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1309
lemma Lim_at_infinity:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1310
  "(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
  1311
  by (auto simp add: tendsto_iff eventually_at_infinity)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1312
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1313
lemma Lim_eventually: "eventually (\<lambda>x. f x = l) net \<Longrightarrow> (f ---> l) net"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1314
  by (rule topological_tendstoI, auto elim: eventually_rev_mono)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1315
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1316
text{* The expected monotonicity property. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1317
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1318
lemma Lim_within_empty: "(f ---> l) (net within {})"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1319
  unfolding tendsto_def Limits.eventually_within by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1320
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1321
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
  1322
  unfolding tendsto_def Limits.eventually_within
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1323
  by (auto elim!: eventually_elim1)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1324
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1325
lemma Lim_Un: assumes "(f ---> l) (net within S)" "(f ---> l) (net within T)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1326
  shows "(f ---> l) (net within (S \<union> T))"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1327
  using assms unfolding tendsto_def Limits.eventually_within
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1328
  apply clarify
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1329
  apply (drule spec, drule (1) mp, drule (1) mp)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1330
  apply (drule spec, drule (1) mp, drule (1) mp)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1331
  apply (auto elim: eventually_elim2)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1332
  done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1333
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1334
lemma Lim_Un_univ:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1335
 "(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
  1336
        ==> (f ---> l) net"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1337
  by (metis Lim_Un within_UNIV)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1338
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1339
text{* Interrelations between restricted and unrestricted limits. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1340
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1341
lemma Lim_at_within: "(f ---> l) net ==> (f ---> l)(net within S)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1342
  (* FIXME: rename *)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1343
  unfolding tendsto_def Limits.eventually_within
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1344
  apply (clarify, drule spec, drule (1) mp, drule (1) mp)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1345
  by (auto elim!: eventually_elim1)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1346
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1347
lemma eventually_within_interior:
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1348
  assumes "x \<in> interior S"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1349
  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
  1350
proof-
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1351
  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
  1352
  { assume "?lhs"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1353
    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
  1354
      unfolding Limits.eventually_within Limits.eventually_at_topological
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1355
      by auto
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1356
    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
  1357
      by auto
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1358
    then have "?rhs"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1359
      unfolding Limits.eventually_at_topological by auto
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1360
  } moreover
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1361
  { assume "?rhs" hence "?lhs"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1362
      unfolding Limits.eventually_within
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1363
      by (auto elim: eventually_elim1)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1364
  } ultimately
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1365
  show "?thesis" ..
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1366
qed
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1367
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1368
lemma at_within_interior:
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1369
  "x \<in> interior S \<Longrightarrow> at x within S = at x"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1370
  by (simp add: filter_eq_iff eventually_within_interior)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1371
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1372
lemma at_within_open:
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1373
  "\<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
  1374
  by (simp only: at_within_interior interior_open)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1375
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1376
lemma Lim_within_open:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1377
  fixes f :: "'a::topological_space \<Rightarrow> 'b::topological_space"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1378
  assumes"a \<in> S" "open S"
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1379
  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
  1380
  using assms by (simp only: at_within_open)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1381
43338
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1382
lemma Lim_within_LIMSEQ:
44584
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  1383
  fixes a :: "'a::metric_space"
43338
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1384
  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
  1385
  shows "(X ---> L) (at a within T)"
44584
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  1386
  using assms unfolding tendsto_def [where l=L]
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  1387
  by (simp add: sequentially_imp_eventually_within)
43338
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1388
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1389
lemma Lim_right_bound:
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1390
  fixes f :: "real \<Rightarrow> real"
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1391
  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
  1392
  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
  1393
  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
  1394
proof cases
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1395
  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
  1396
next
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1397
  assume [simp]: "{x<..} \<inter> I \<noteq> {}"
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1398
  show ?thesis
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1399
  proof (rule Lim_within_LIMSEQ, safe)
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1400
    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
  1401
    
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1402
    show "(\<lambda>n. f (S n)) ----> Inf (f ` ({x<..} \<inter> I))"
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1403
    proof (rule LIMSEQ_I, rule ccontr)
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1404
      fix r :: real assume "0 < r"
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1405
      with Inf_close[of "f ` ({x<..} \<inter> I)" r]
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1406
      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
  1407
      from `x < y` have "0 < y - x" by auto
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1408
      from S(2)[THEN LIMSEQ_D, OF this]
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1409
      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
  1410
      
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1411
      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
  1412
      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
  1413
        using S bnd by (intro Inf_lower[where z=K]) auto
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1414
      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
  1415
        by (auto simp: not_less field_simps)
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1416
      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
  1417
      show False by auto
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1418
    qed
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1419
  qed
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1420
qed
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1421
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1422
text{* Another limit point characterization. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1423
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1424
lemma islimpt_sequential:
50883
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1425
  fixes x :: "'a::first_countable_topology"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1426
  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
  1427
    (is "?lhs = ?rhs")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1428
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1429
  assume ?lhs
50883
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1430
  from countable_basis_at_decseq[of x] guess A . note A = this
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1431
  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
  1432
  { fix n
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1433
    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
  1434
      unfolding islimpt_def using A(1,2)[of n] by auto
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1435
    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
  1436
      unfolding f_def by (rule someI_ex)
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1437
    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
  1438
  then have "\<forall>n. f n \<in> S - {x}" by auto
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1439
  moreover have "(\<lambda>n. f n) ----> x"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1440
  proof (rule topological_tendstoI)
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1441
    fix S assume "open S" "x \<in> S"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1442
    from A(3)[OF this] `\<And>n. f n \<in> A n`
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1443
    show "eventually (\<lambda>x. f x \<in> S) sequentially" by (auto elim!: eventually_elim1)
44584
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  1444
  qed
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  1445
  ultimately show ?rhs by fast
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1446
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1447
  assume ?rhs
50883
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1448
  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
  1449
  show ?lhs
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1450
    unfolding islimpt_def
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1451
  proof safe
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1452
    fix T assume "open T" "x \<in> T"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1453
    from lim[THEN topological_tendstoD, OF this] f
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1454
    show "\<exists>y\<in>S. y \<in> T \<and> y \<noteq> x"
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1455
      unfolding eventually_sequentially by auto
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1456
  qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1457
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1458
44125
230a8665c919 mark some redundant theorems as legacy
huffman
parents: 44122
diff changeset
  1459
lemma Lim_inv: (* TODO: delete *)
44081
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44076
diff changeset
  1460
  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
  1461
  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
  1462
  shows "((inverse o f) ---> inverse l) A"
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  1463
  unfolding o_def using assms by (rule tendsto_inverse)
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  1464
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1465
lemma Lim_null:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1466
  fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
44125
230a8665c919 mark some redundant theorems as legacy
huffman
parents: 44122
diff changeset
  1467
  shows "(f ---> l) net \<longleftrightarrow> ((\<lambda>x. f(x) - l) ---> 0) net"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1468
  by (simp add: Lim dist_norm)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1469
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1470
lemma Lim_null_comparison:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1471
  fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1472
  assumes "eventually (\<lambda>x. norm (f x) \<le> g x) net" "(g ---> 0) net"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1473
  shows "(f ---> 0) net"
44252
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1474
proof (rule metric_tendsto_imp_tendsto)
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1475
  show "(g ---> 0) net" by fact
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1476
  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
  1477
    using assms(1) by (rule eventually_elim1, simp add: dist_norm)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1478
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1479
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1480
lemma Lim_transform_bound:
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
  fixes g :: "'a \<Rightarrow> 'c::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1483
  assumes "eventually (\<lambda>n. norm(f n) <= norm(g n)) net"  "(g ---> 0) net"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1484
  shows "(f ---> 0) net"
44252
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1485
  using assms(1) tendsto_norm_zero [OF assms(2)]
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1486
  by (rule Lim_null_comparison)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1487
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1488
text{* Deducing things about the limit from the elements. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1489
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1490
lemma Lim_in_closed_set:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1491
  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
  1492
  shows "l \<in> S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1493
proof (rule ccontr)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1494
  assume "l \<notin> S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1495
  with `closed S` have "open (- S)" "l \<in> - S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1496
    by (simp_all add: open_Compl)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1497
  with assms(4) have "eventually (\<lambda>x. f x \<in> - S) net"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1498
    by (rule topological_tendstoD)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1499
  with assms(2) have "eventually (\<lambda>x. False) net"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1500
    by (rule eventually_elim2) simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1501
  with assms(3) show "False"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1502
    by (simp add: eventually_False)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1503
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1504
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1505
text{* Need to prove closed(cball(x,e)) before deducing this as a corollary. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1506
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1507
lemma Lim_dist_ubound:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1508
  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
  1509
  shows "dist a l <= e"
44252
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1510
proof-
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1511
  have "dist a l \<in> {..e}"
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1512
  proof (rule Lim_in_closed_set)
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1513
    show "closed {..e}" by simp
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1514
    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
  1515
    show "\<not> trivial_limit net" by fact
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1516
    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
  1517
  qed
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1518
  thus ?thesis by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1519
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1520
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1521
lemma Lim_norm_ubound:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1522
  fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1523
  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
  1524
  shows "norm(l) <= e"
44252
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1525
proof-
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1526
  have "norm l \<in> {..e}"
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1527
  proof (rule Lim_in_closed_set)
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1528
    show "closed {..e}" by simp
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1529
    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
  1530
    show "\<not> trivial_limit net" by fact
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1531
    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
  1532
  qed
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1533
  thus ?thesis by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1534
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1535
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1536
lemma Lim_norm_lbound:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1537
  fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1538
  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
  1539
  shows "e \<le> norm l"
44252
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1540
proof-
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1541
  have "norm l \<in> {e..}"
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1542
  proof (rule Lim_in_closed_set)
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1543
    show "closed {e..}" by simp
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1544
    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
  1545
    show "\<not> trivial_limit net" by fact
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1546
    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
  1547
  qed
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1548
  thus ?thesis by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1549
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1550
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1551
text{* Uniqueness of the limit, when nontrivial. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1552
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1553
lemma tendsto_Lim:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1554
  fixes f :: "'a \<Rightarrow> 'b::t2_space"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1555
  shows "~(trivial_limit net) \<Longrightarrow> (f ---> l) net ==> Lim net f = l"
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41969
diff changeset
  1556
  unfolding Lim_def using tendsto_unique[of net f] by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1557
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1558
text{* Limit under bilinear function *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1559
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1560
lemma Lim_bilinear:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1561
  assumes "(f ---> l) net" and "(g ---> m) net" and "bounded_bilinear h"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1562
  shows "((\<lambda>x. h (f x) (g x)) ---> (h l m)) net"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1563
using `bounded_bilinear h` `(f ---> l) net` `(g ---> m) net`
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1564
by (rule bounded_bilinear.tendsto)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1565
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1566
text{* These are special for limits out of the same vector space. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1567
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1568
lemma Lim_within_id: "(id ---> a) (at a within s)"
45031
9583f2b56f85 add lemmas within_empty and tendsto_bot;
huffman
parents: 44909
diff changeset
  1569
  unfolding id_def by (rule tendsto_ident_at_within)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1570
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1571
lemma Lim_at_id: "(id ---> a) (at a)"
45031
9583f2b56f85 add lemmas within_empty and tendsto_bot;
huffman
parents: 44909
diff changeset
  1572
  unfolding id_def by (rule tendsto_ident_at)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1573
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1574
lemma Lim_at_zero:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1575
  fixes a :: "'a::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1576
  fixes l :: "'b::topological_space"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1577
  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
  1578
  using LIM_offset_zero LIM_offset_zero_cancel ..
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1579
44081
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44076
diff changeset
  1580
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
  1581
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1582
definition
44081
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44076
diff changeset
  1583
  netlimit :: "'a::t2_space filter \<Rightarrow> 'a" where
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1584
  "netlimit net = (SOME a. ((\<lambda>x. x) ---> a) net)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1585
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1586
lemma netlimit_within:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1587
  assumes "\<not> trivial_limit (at a within S)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1588
  shows "netlimit (at a within S) = a"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1589
unfolding netlimit_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1590
apply (rule some_equality)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1591
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
  1592
apply (rule tendsto_ident_at)
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41969
diff changeset
  1593
apply (erule tendsto_unique [OF assms])
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1594
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
  1595
apply (rule tendsto_ident_at)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1596
done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1597
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1598
lemma netlimit_at:
44072
5b970711fb39 class perfect_space inherits from topological_space;
huffman
parents: 43338
diff changeset
  1599
  fixes a :: "'a::{perfect_space,t2_space}"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1600
  shows "netlimit (at a) = a"
45031
9583f2b56f85 add lemmas within_empty and tendsto_bot;
huffman
parents: 44909
diff changeset
  1601
  using netlimit_within [of a UNIV] by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1602
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1603
lemma lim_within_interior:
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1604
  "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
  1605
  by (simp add: at_within_interior)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1606
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1607
lemma netlimit_within_interior:
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1608
  fixes x :: "'a::{t2_space,perfect_space}"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1609
  assumes "x \<in> interior S"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1610
  shows "netlimit (at x within S) = x"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1611
using assms by (simp add: at_within_interior netlimit_at)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1612
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1613
text{* Transformation of limit. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1614
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1615
lemma Lim_transform:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1616
  fixes f g :: "'a::type \<Rightarrow> 'b::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1617
  assumes "((\<lambda>x. f x - g x) ---> 0) net" "(f ---> l) net"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1618
  shows "(g ---> l) net"
44252
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1619
  using tendsto_diff [OF assms(2) assms(1)] by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1620
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1621
lemma Lim_transform_eventually:
36667
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1622
  "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
  1623
  apply (rule topological_tendstoI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1624
  apply (drule (2) topological_tendstoD)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1625
  apply (erule (1) eventually_elim2, simp)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1626
  done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1627
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1628
lemma Lim_transform_within:
36667
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1629
  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
  1630
  and "(f ---> l) (at x within S)"
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1631
  shows "(g ---> l) (at x within S)"
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1632
proof (rule Lim_transform_eventually)
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1633
  show "eventually (\<lambda>x. f x = g x) (at x within S)"
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1634
    unfolding eventually_within
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1635
    using assms(1,2) by auto
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1636
  show "(f ---> l) (at x within S)" by fact
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1637
qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1638
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1639
lemma Lim_transform_at:
36667
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1640
  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
  1641
  and "(f ---> l) (at x)"
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1642
  shows "(g ---> l) (at x)"
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1643
proof (rule Lim_transform_eventually)
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1644
  show "eventually (\<lambda>x. f x = g x) (at x)"
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1645
    unfolding eventually_at
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1646
    using assms(1,2) by auto
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1647
  show "(f ---> l) (at x)" by fact
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1648
qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1649
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1650
text{* Common case assuming being away from some crucial point like 0. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1651
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1652
lemma Lim_transform_away_within:
36669
c90c8a3ae1f7 generalize some lemmas to class t1_space
huffman
parents: 36668
diff changeset
  1653
  fixes a b :: "'a::t1_space"
36667
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1654
  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
  1655
  and "(f ---> l) (at a within S)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1656
  shows "(g ---> l) (at a within S)"
36669
c90c8a3ae1f7 generalize some lemmas to class t1_space
huffman
parents: 36668
diff changeset
  1657
proof (rule Lim_transform_eventually)
c90c8a3ae1f7 generalize some lemmas to class t1_space
huffman
parents: 36668
diff changeset
  1658
  show "(f ---> l) (at a within S)" by fact
c90c8a3ae1f7 generalize some lemmas to class t1_space
huffman
parents: 36668
diff changeset
  1659
  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
  1660
    unfolding Limits.eventually_within eventually_at_topological
c90c8a3ae1f7 generalize some lemmas to class t1_space
huffman
parents: 36668
diff changeset
  1661
    by (rule exI [where x="- {b}"], simp add: open_Compl assms)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1662
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1663
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1664
lemma Lim_transform_away_at:
36669
c90c8a3ae1f7 generalize some lemmas to class t1_space
huffman
parents: 36668
diff changeset
  1665
  fixes a b :: "'a::t1_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1666
  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
  1667
  and fl: "(f ---> l) (at a)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1668
  shows "(g ---> l) (at a)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1669
  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
  1670
  by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1671
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1672
text{* Alternatively, within an open set. *}
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_within_open:
36667
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1675
  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
  1676
  and "(f ---> l) (at a)"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1677
  shows "(g ---> l) (at a)"
36667
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1678
proof (rule Lim_transform_eventually)
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1679
  show "eventually (\<lambda>x. f x = g x) (at a)"
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1680
    unfolding eventually_at_topological
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1681
    using assms(1,2,3) by auto
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1682
  show "(f ---> l) (at a)" by fact
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1683
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1684
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1685
text{* A congruence rule allowing us to transform limits assuming not at point. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1686
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1687
(* FIXME: Only one congruence rule for tendsto can be used at a time! *)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1688
36362
06475a1547cb fix lots of looping simp calls and other warnings
huffman
parents: 36360
diff changeset
  1689
lemma Lim_cong_within(*[cong add]*):
43338
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1690
  assumes "a = b" "x = y" "S = T"
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1691
  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
  1692
  shows "(f ---> x) (at a within S) \<longleftrightarrow> (g ---> y) (at b within T)"
36667
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1693
  unfolding tendsto_def Limits.eventually_within eventually_at_topological
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1694
  using assms by simp
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1695
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1696
lemma Lim_cong_at(*[cong add]*):
43338
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1697
  assumes "a = b" "x = y"
36667
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1698
  assumes "\<And>x. x \<noteq> a \<Longrightarrow> f x = g x"
43338
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1699
  shows "((\<lambda>x. f x) ---> x) (at a) \<longleftrightarrow> ((g ---> y) (at a))"
36667
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1700
  unfolding tendsto_def eventually_at_topological
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1701
  using assms by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1702
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1703
text{* Useful lemmas on closure and set of possible sequential limits.*}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1704
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1705
lemma closure_sequential:
50883
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1706
  fixes l :: "'a::first_countable_topology"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1707
  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
  1708
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1709
  assume "?lhs" moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1710
  { assume "l \<in> S"
44125
230a8665c919 mark some redundant theorems as legacy
huffman
parents: 44122
diff changeset
  1711
    hence "?rhs" using tendsto_const[of l sequentially] by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1712
  } moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1713
  { assume "l islimpt S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1714
    hence "?rhs" unfolding islimpt_sequential by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1715
  } ultimately
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1716
  show "?rhs" unfolding closure_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1717
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1718
  assume "?rhs"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1719
  thus "?lhs" unfolding closure_def unfolding islimpt_sequential by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1720
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1721
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1722
lemma closed_sequential_limits:
50883
1421884baf5b introduce first_countable_topology typeclass
hoelzl
parents: 50882
diff changeset
  1723
  fixes S :: "'a::first_countable_topology set"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1724
  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
  1725
  unfolding closed_limpt
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1726
  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
  1727
  by metis
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1728
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1729
lemma closure_approachable:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1730
  fixes S :: "'a::metric_space set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1731
  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
  1732
  apply (auto simp add: closure_def islimpt_approachable)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1733
  by (metis dist_self)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1734
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1735
lemma closed_approachable:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1736
  fixes S :: "'a::metric_space set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1737
  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
  1738
  by (metis closure_closed closure_approachable)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1739
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1740
subsection {* Infimum Distance *}
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1741
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1742
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
  1743
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1744
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
  1745
  by (simp add: infdist_def)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1746
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1747
lemma infdist_nonneg:
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1748
  shows "0 \<le> infdist x A"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1749
  using assms by (auto simp add: infdist_def)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1750
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1751
lemma infdist_le:
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1752
  assumes "a \<in> A"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1753
  assumes "d = dist x a"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1754
  shows "infdist x A \<le> d"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1755
  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
  1756
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1757
lemma infdist_zero[simp]:
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1758
  assumes "a \<in> A" shows "infdist a A = 0"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1759
proof -
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1760
  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
  1761
  with infdist_nonneg[of a A] assms show "infdist a A = 0" by auto
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1762
qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1763
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1764
lemma infdist_triangle:
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1765
  shows "infdist x A \<le> infdist y A + dist x y"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1766
proof cases
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1767
  assume "A = {}" thus ?thesis by (simp add: infdist_def)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1768
next
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1769
  assume "A \<noteq> {}" then obtain a where "a \<in> A" by auto
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1770
  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
  1771
  proof
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1772
    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
  1773
    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
  1774
    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
  1775
    show "infdist x A \<le> d"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1776
      unfolding infdist_notempty[OF `A \<noteq> {}`]
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1777
    proof (rule Inf_lower2)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1778
      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
  1779
      show "dist x a \<le> d" unfolding d by (rule dist_triangle)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1780
      fix d assume "d \<in> {dist x a |a. a \<in> A}"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1781
      then obtain a where "a \<in> A" "d = dist x a" by auto
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1782
      thus "infdist x A \<le> d" by (rule infdist_le)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1783
    qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1784
  qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1785
  also have "\<dots> = dist x y + infdist y A"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1786
  proof (rule Inf_eq, safe)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1787
    fix a assume "a \<in> A"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1788
    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
  1789
  next
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1790
    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
  1791
    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
  1792
      by (intro Inf_greatest) (auto simp: field_simps)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1793
    thus "i \<le> dist x y + infdist y A" by simp
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1794
  qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1795
  finally show ?thesis by simp
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1796
qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1797
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1798
lemma
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1799
  in_closure_iff_infdist_zero:
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1800
  assumes "A \<noteq> {}"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1801
  shows "x \<in> closure A \<longleftrightarrow> infdist x A = 0"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1802
proof
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1803
  assume "x \<in> closure A"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1804
  show "infdist x A = 0"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1805
  proof (rule ccontr)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1806
    assume "infdist x A \<noteq> 0"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1807
    with infdist_nonneg[of x A] have "infdist x A > 0" by auto
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1808
    hence "ball x (infdist x A) \<inter> closure A = {}" apply auto
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1809
      by (metis `0 < infdist x A` `x \<in> closure A` closure_approachable dist_commute
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1810
        eucl_less_not_refl euclidean_trans(2) infdist_le)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1811
    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
  1812
    thus False using `x \<in> closure A` by simp
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1813
  qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1814
next
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1815
  assume x: "infdist x A = 0"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1816
  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
  1817
  show "x \<in> closure A" unfolding closure_approachable
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1818
  proof (safe, rule ccontr)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1819
    fix e::real assume "0 < e"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1820
    assume "\<not> (\<exists>y\<in>A. dist y x < e)"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1821
    hence "infdist x A \<ge> e" using `a \<in> A`
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1822
      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
  1823
      by (force simp: dist_commute)
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1824
    with x `0 < e` show False by auto
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1825
  qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1826
qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1827
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1828
lemma
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1829
  in_closed_iff_infdist_zero:
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1830
  assumes "closed A" "A \<noteq> {}"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1831
  shows "x \<in> A \<longleftrightarrow> infdist x A = 0"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1832
proof -
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1833
  have "x \<in> closure A \<longleftrightarrow> infdist x A = 0"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1834
    by (rule in_closure_iff_infdist_zero) fact
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1835
  with assms show ?thesis by simp
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 tendsto_infdist [tendsto_intros]:
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1839
  assumes f: "(f ---> l) F"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1840
  shows "((\<lambda>x. infdist (f x) A) ---> infdist l A) F"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1841
proof (rule tendstoI)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1842
  fix e ::real assume "0 < e"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1843
  from tendstoD[OF f this]
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1844
  show "eventually (\<lambda>x. dist (infdist (f x) A) (infdist l A) < e) F"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1845
  proof (eventually_elim)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1846
    fix x
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1847
    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
  1848
    have "dist (infdist (f x) A) (infdist l A) \<le> dist (f x) l"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1849
      by (simp add: dist_commute dist_real_def)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1850
    also assume "dist (f x) l < e"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1851
    finally show "dist (infdist (f x) A) (infdist l A) < e" .
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1852
  qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1853
qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1854
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1855
text{* Some other lemmas about sequences. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1856
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  1857
lemma sequentially_offset:
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  1858
  assumes "eventually (\<lambda>i. P i) sequentially"
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  1859
  shows "eventually (\<lambda>i. P (i + k)) sequentially"
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  1860
  using assms unfolding eventually_sequentially by (metis trans_le_add1)
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  1861
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1862
lemma seq_offset:
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  1863
  assumes "(f ---> l) sequentially"
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  1864
  shows "((\<lambda>i. f (i + k)) ---> l) sequentially"
44584
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  1865
  using assms by (rule LIMSEQ_ignore_initial_segment) (* FIXME: redundant *)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1866
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1867
lemma seq_offset_neg:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1868
  "(f ---> l) sequentially ==> ((\<lambda>i. f(i - k)) ---> l) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1869
  apply (rule topological_tendstoI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1870
  apply (drule (2) topological_tendstoD)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1871
  apply (simp only: eventually_sequentially)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1872
  apply (subgoal_tac "\<And>N k (n::nat). N + k <= n ==> N <= n - k")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1873
  apply metis
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1874
  by arith
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1875
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1876
lemma seq_offset_rev:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1877
  "((\<lambda>i. f(i + k)) ---> l) sequentially ==> (f ---> l) sequentially"
44584
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  1878
  by (rule LIMSEQ_offset) (* FIXME: redundant *)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1879
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1880
lemma seq_harmonic: "((\<lambda>n. inverse (real n)) ---> 0) sequentially"
44584
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  1881
  using LIMSEQ_inverse_real_of_nat by (rule LIMSEQ_imp_Suc)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1882
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1883
subsection {* More properties of closed balls *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1884
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1885
lemma closed_cball: "closed (cball x e)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1886
unfolding cball_def closed_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1887
unfolding Collect_neg_eq [symmetric] not_le
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1888
apply (clarsimp simp add: open_dist, rename_tac y)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1889
apply (rule_tac x="dist x y - e" in exI, clarsimp)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1890
apply (rename_tac x')
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1891
apply (cut_tac x=x and y=x' and z=y in dist_triangle)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1892
apply simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1893
done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1894
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1895
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
  1896
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1897
  { 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
  1898
    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
  1899
  } moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1900
  { 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
  1901
    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
  1902
  } ultimately
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1903
  show ?thesis unfolding open_contains_ball by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1904
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1905
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1906
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
  1907
  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
  1908
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1909
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
  1910
  apply (simp add: interior_def, safe)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1911
  apply (force simp add: open_contains_cball)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1912
  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
  1913
  apply (simp add: subset_trans [OF ball_subset_cball])
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1914
  done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1915
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1916
lemma islimpt_ball:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1917
  fixes x y :: "'a::{real_normed_vector,perfect_space}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1918
  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
  1919
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1920
  assume "?lhs"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1921
  { assume "e \<le> 0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1922
    hence *:"ball x e = {}" using ball_eq_empty[of x e] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1923
    have False using `?lhs` unfolding * using islimpt_EMPTY[of y] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1924
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1925
  hence "e > 0" by (metis not_less)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1926
  moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1927
  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
  1928
  ultimately show "?rhs" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1929
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1930
  assume "?rhs" hence "e>0"  by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1931
  { fix d::real assume "d>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1932
    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
  1933
    proof(cases "d \<le> dist x y")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1934
      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
  1935
      proof(cases "x=y")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1936
        case True hence False using `d \<le> dist x y` `d>0` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1937
        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
  1938
      next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1939
        case False
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1940
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1941
        have "dist x (y - (d / (2 * dist y x)) *\<^sub>R (y - x))
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1942
              = norm (x - y + (d / (2 * norm (y - x))) *\<^sub>R (y - x))"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1943
          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
  1944
        also have "\<dots> = \<bar>- 1 + d / (2 * norm (x - y))\<bar> * norm (x - y)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1945
          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
  1946
          unfolding scaleR_minus_left scaleR_one
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1947
          by (auto simp add: norm_minus_commute)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1948
        also have "\<dots> = \<bar>- norm (x - y) + d / 2\<bar>"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1949
          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
  1950
          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
  1951
        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
  1952
        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
  1953
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1954
        moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1955
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1956
        have "(d / (2*dist y x)) *\<^sub>R (y - x) \<noteq> 0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1957
          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
  1958
        moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1959
        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
  1960
          using `d>0` `x\<noteq>y`[unfolded dist_nz] dist_commute[of x y]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1961
          unfolding dist_norm by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1962
        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
  1963
      qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1964
    next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1965
      case False hence "d > dist x y" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1966
      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
  1967
      proof(cases "x=y")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1968
        case True
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1969
        obtain z where **: "z \<noteq> y" "dist z y < min e d"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1970
          using perfect_choose_dist[of "min e d" y]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1971
          using `d > 0` `e>0` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1972
        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
  1973
          unfolding `x = y`
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1974
          using `z \<noteq> y` **
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1975
          by (rule_tac x=z in bexI, auto simp add: dist_commute)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1976
      next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1977
        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
  1978
          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
  1979
      qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1980
    qed  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1981
  thus "?lhs" unfolding mem_cball islimpt_approachable mem_ball by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1982
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1983
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1984
lemma closure_ball_lemma:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1985
  fixes x y :: "'a::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1986
  assumes "x \<noteq> y" shows "y islimpt ball x (dist x y)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1987
proof (rule islimptI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1988
  fix T assume "y \<in> T" "open T"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1989
  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
  1990
    unfolding open_dist by fast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1991
  (* choose point between x and y, within distance r of y. *)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1992
  def k \<equiv> "min 1 (r / (2 * dist x y))"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1993
  def z \<equiv> "y + scaleR k (x - y)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1994
  have z_def2: "z = x + scaleR (1 - k) (y - x)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1995
    unfolding z_def by (simp add: algebra_simps)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1996
  have "dist z y < r"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1997
    unfolding z_def k_def using `0 < r`
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1998
    by (simp add: dist_norm min_def)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1999
  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
  2000
  have "dist x z < dist x y"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2001
    unfolding z_def2 dist_norm
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2002
    apply (simp add: norm_minus_commute)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2003
    apply (simp only: dist_norm [symmetric])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2004
    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
  2005
    apply (rule mult_strict_right_mono)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2006
    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
  2007
    apply (simp add: zero_less_dist_iff `x \<noteq> y`)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2008
    done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2009
  hence "z \<in> ball x (dist x y)" by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2010
  have "z \<noteq> y"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2011
    unfolding z_def k_def using `x \<noteq> y` `0 < r`
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2012
    by (simp add: min_def)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2013
  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
  2014
    using `z \<in> ball x (dist x y)` `z \<in> T` `z \<noteq> y`
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2015
    by fast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2016
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2017
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2018
lemma closure_ball:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2019
  fixes x :: "'a::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2020
  shows "0 < e \<Longrightarrow> closure (ball x e) = cball x e"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2021
apply (rule equalityI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2022
apply (rule closure_minimal)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2023
apply (rule ball_subset_cball)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2024
apply (rule closed_cball)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2025
apply (rule subsetI, rename_tac y)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2026
apply (simp add: le_less [where 'a=real])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2027
apply (erule disjE)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2028
apply (rule subsetD [OF closure_subset], simp)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2029
apply (simp add: closure_def)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2030
apply clarify
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2031
apply (rule closure_ball_lemma)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2032
apply (simp add: zero_less_dist_iff)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2033
done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2034
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2035
(* In a trivial vector space, this fails for e = 0. *)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2036
lemma interior_cball:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2037
  fixes x :: "'a::{real_normed_vector, perfect_space}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2038
  shows "interior (cball x e) = ball x e"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2039
proof(cases "e\<ge>0")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2040
  case False note cs = this
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2041
  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
  2042
  { fix y assume "y \<in> cball x e"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2043
    hence False unfolding mem_cball using dist_nz[of x y] cs by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2044
  hence "cball x e = {}" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2045
  hence "interior (cball x e) = {}" using interior_empty by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2046
  ultimately show ?thesis by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2047
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2048
  case True note cs = this
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2049
  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
  2050
  { 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
  2051
    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
  2052
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2053
    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
  2054
      using perfect_choose_dist [of d] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2055
    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
  2056
    hence xa_cball:"xa \<in> cball x e" using as(1) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2057
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2058
    hence "y \<in> ball x e" proof(cases "x = y")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2059
      case True
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2060
      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
  2061
      thus "y \<in> ball x e" using `x = y ` by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2062
    next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2063
      case False
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2064
      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
  2065
        using `d>0` norm_ge_zero[of "y - x"] `x \<noteq> y` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2066
      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
  2067
      have "y - x \<noteq> 0" using `x \<noteq> y` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2068
      hence **:"d / (2 * norm (y - x)) > 0" unfolding zero_less_norm_iff[THEN sym]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2069
        using `d>0` divide_pos_pos[of d "2*norm (y - x)"] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2070
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2071
      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
  2072
        by (auto simp add: dist_norm algebra_simps)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2073
      also have "\<dots> = norm ((1 + d / (2 * norm (y - x))) *\<^sub>R (y - x))"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2074
        by (auto simp add: algebra_simps)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2075
      also have "\<dots> = \<bar>1 + d / (2 * norm (y - x))\<bar> * norm (y - x)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2076
        using ** by auto
49962
a8cc904a6820 Renamed {left,right}_distrib to distrib_{right,left}.
webertj
parents: 49834
diff changeset
  2077
      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
  2078
      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
  2079
      thus "y \<in> ball x e" unfolding mem_ball using `d>0` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2080
    qed  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2081
  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
  2082
  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
  2083
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2084
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2085
lemma frontier_ball:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2086
  fixes a :: "'a::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2087
  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
  2088
  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
  2089
  apply (simp add: set_eq_iff)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2090
  by arith
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2091
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2092
lemma frontier_cball:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2093
  fixes a :: "'a::{real_normed_vector, perfect_space}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2094
  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
  2095
  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
  2096
  apply (simp add: set_eq_iff)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2097
  by arith
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2098
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2099
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
  2100
  apply (simp add: set_eq_iff not_le)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2101
  by (metis zero_le_dist dist_self order_less_le_trans)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2102
lemma cball_empty: "e < 0 ==> cball x e = {}" by (simp add: cball_eq_empty)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2103
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2104
lemma cball_eq_sing:
44072
5b970711fb39 class perfect_space inherits from topological_space;
huffman
parents: 43338
diff changeset
  2105
  fixes x :: "'a::{metric_space,perfect_space}"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2106
  shows "(cball x e = {x}) \<longleftrightarrow> e = 0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2107
proof (rule linorder_cases)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2108
  assume e: "0 < e"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2109
  obtain a where "a \<noteq> x" "dist a x < e"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2110
    using perfect_choose_dist [OF e] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2111
  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
  2112
  with e show ?thesis by (auto simp add: set_eq_iff)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2113
qed auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2114
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2115
lemma cball_sing:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2116
  fixes x :: "'a::metric_space"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2117
  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
  2118
  by (auto simp add: set_eq_iff)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2119
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  2120
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  2121
subsection {* Boundedness *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2122
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2123
  (* 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
  2124
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
  2125
  bounded :: "'a set \<Rightarrow> bool" where
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2126
  "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
  2127
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2128
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
  2129
unfolding bounded_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2130
apply safe
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2131
apply (rule_tac x="dist a x + e" in exI, clarify)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2132
apply (drule (1) bspec)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2133
apply (erule order_trans [OF dist_triangle add_left_mono])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2134
apply auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2135
done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2136
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2137
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
  2138
unfolding bounded_any_center [where a=0]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2139
by (simp add: dist_norm)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2140
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50094
diff changeset
  2141
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
  2142
  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
  2143
  using assms by auto
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50094
diff changeset
  2144
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2145
lemma bounded_empty[simp]: "bounded {}" by (simp add: bounded_def)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2146
lemma bounded_subset: "bounded T \<Longrightarrow> S \<subseteq> T ==> bounded S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2147
  by (metis bounded_def subset_eq)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2148
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2149
lemma bounded_interior[intro]: "bounded S ==> bounded(interior S)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2150
  by (metis bounded_subset interior_subset)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2151
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2152
lemma bounded_closure[intro]: assumes "bounded S" shows "bounded(closure S)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2153
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2154
  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
  2155
  { fix y assume "y \<in> closure S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2156
    then obtain f where f: "\<forall>n. f n \<in> S"  "(f ---> y) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2157
      unfolding closure_sequential by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2158
    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
  2159
    hence "eventually (\<lambda>n. dist x (f n) \<le> a) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2160
      by (rule eventually_mono, simp add: f(1))
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2161
    have "dist x y \<le> a"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2162
      apply (rule Lim_dist_ubound [of sequentially f])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2163
      apply (rule trivial_limit_sequentially)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2164
      apply (rule f(2))
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2165
      apply fact
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2166
      done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2167
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2168
  thus ?thesis unfolding bounded_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2169
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2170
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2171
lemma bounded_cball[simp,intro]: "bounded (cball x e)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2172
  apply (simp add: bounded_def)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2173
  apply (rule_tac x=x in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2174
  apply (rule_tac x=e in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2175
  apply auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2176
  done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2177
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2178
lemma bounded_ball[simp,intro]: "bounded(ball x e)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2179
  by (metis ball_subset_cball bounded_cball bounded_subset)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2180
36362
06475a1547cb fix lots of looping simp calls and other warnings
huffman
parents: 36360
diff changeset
  2181
lemma finite_imp_bounded[intro]:
06475a1547cb fix lots of looping simp calls and other warnings
huffman
parents: 36360
diff changeset
  2182
  fixes S :: "'a::metric_space set" assumes "finite S" shows "bounded S"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2183
proof-
36362
06475a1547cb fix lots of looping simp calls and other warnings
huffman
parents: 36360
diff changeset
  2184
  { fix a and F :: "'a set" assume as:"bounded F"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2185
    then obtain x e where "\<forall>y\<in>F. dist x y \<le> e" unfolding bounded_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2186
    hence "\<forall>y\<in>(insert a F). dist x y \<le> max e (dist x a)" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2187
    hence "bounded (insert a F)" unfolding bounded_def by (intro exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2188
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2189
  thus ?thesis using finite_induct[of S bounded]  using bounded_empty assms by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2190
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2191
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2192
lemma bounded_Un[simp]: "bounded (S \<union> T) \<longleftrightarrow> bounded S \<and> bounded T"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2193
  apply (auto simp add: bounded_def)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2194
  apply (rename_tac x y r s)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2195
  apply (rule_tac x=x in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2196
  apply (rule_tac x="max r (dist x y + s)" in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2197
  apply (rule ballI, rename_tac z, safe)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2198
  apply (drule (1) bspec, simp)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2199
  apply (drule (1) bspec)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2200
  apply (rule min_max.le_supI2)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2201
  apply (erule order_trans [OF dist_triangle add_left_mono])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2202
  done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2203
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2204
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
  2205
  by (induct rule: finite_induct[of F], auto)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2206
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2207
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
  2208
  apply (simp add: bounded_iff)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2209
  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
  2210
  by metis arith
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2211
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2212
lemma bounded_Int[intro]: "bounded S \<or> bounded T \<Longrightarrow> bounded (S \<inter> T)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2213
  by (metis Int_lower1 Int_lower2 bounded_subset)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2214
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2215
lemma bounded_diff[intro]: "bounded S ==> bounded (S - T)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2216
apply (metis Diff_subset bounded_subset)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2217
done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2218
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2219
lemma bounded_insert[intro]:"bounded(insert x S) \<longleftrightarrow> bounded S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2220
  by (metis Diff_cancel Un_empty_right Un_insert_right bounded_Un bounded_subset finite.emptyI finite_imp_bounded infinite_remove subset_insertI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2221
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2222
lemma not_bounded_UNIV[simp, intro]:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2223
  "\<not> bounded (UNIV :: 'a::{real_normed_vector, perfect_space} set)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2224
proof(auto simp add: bounded_pos not_le)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2225
  obtain x :: 'a where "x \<noteq> 0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2226
    using perfect_choose_dist [OF zero_less_one] by fast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2227
  fix b::real  assume b: "b >0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2228
  have b1: "b +1 \<ge> 0" using b by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2229
  with `x \<noteq> 0` have "b < norm (scaleR (b + 1) (sgn x))"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2230
    by (simp add: norm_sgn)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2231
  then show "\<exists>x::'a. b < norm x" ..
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2232
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2233
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2234
lemma bounded_linear_image:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2235
  assumes "bounded S" "bounded_linear f"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2236
  shows "bounded(f ` S)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2237
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2238
  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
  2239
  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
  2240
  { fix x assume "x\<in>S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2241
    hence "norm x \<le> b" using b by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2242
    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
  2243
      by (metis B(1) B(2) order_trans mult_le_cancel_left_pos)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2244
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2245
  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
  2246
    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
  2247
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2248
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2249
lemma bounded_scaling:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2250
  fixes S :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2251
  shows "bounded S \<Longrightarrow> bounded ((\<lambda>x. c *\<^sub>R x) ` S)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2252
  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
  2253
  apply (rule bounded_linear_scaleR_right)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2254
  done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2255
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2256
lemma bounded_translation:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2257
  fixes S :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2258
  assumes "bounded S" shows "bounded ((\<lambda>x. a + x) ` S)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2259
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2260
  from assms obtain b where b:"b>0" "\<forall>x\<in>S. norm x \<le> b" unfolding bounded_pos by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2261
  { fix x assume "x\<in>S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2262
    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
  2263
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2264
  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
  2265
    by (auto intro!: exI[of _ "b + norm a"])
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2266
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2267
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2268
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2269
text{* Some theorems on sups and infs using the notion "bounded". *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2270
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2271
lemma bounded_real:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2272
  fixes S :: "real set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2273
  shows "bounded S \<longleftrightarrow>  (\<exists>a. \<forall>x\<in>S. abs x <= a)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2274
  by (simp add: bounded_iff)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2275
33270
paulson
parents: 33175
diff changeset
  2276
lemma bounded_has_Sup:
paulson
parents: 33175
diff changeset
  2277
  fixes S :: "real set"
paulson
parents: 33175
diff changeset
  2278
  assumes "bounded S" "S \<noteq> {}"
paulson
parents: 33175
diff changeset
  2279
  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
  2280
proof
paulson
parents: 33175
diff changeset
  2281
  fix x assume "x\<in>S"
paulson
parents: 33175
diff changeset
  2282
  thus "x \<le> Sup S"
paulson
parents: 33175
diff changeset
  2283
    by (metis SupInf.Sup_upper abs_le_D1 assms(1) bounded_real)
paulson
parents: 33175
diff changeset
  2284
next
paulson
parents: 33175
diff changeset
  2285
  show "\<forall>b. (\<forall>x\<in>S. x \<le> b) \<longrightarrow> Sup S \<le> b" using assms
paulson
parents: 33175
diff changeset
  2286
    by (metis SupInf.Sup_least)
paulson
parents: 33175
diff changeset
  2287
qed
paulson
parents: 33175
diff changeset
  2288
paulson
parents: 33175
diff changeset
  2289
lemma Sup_insert:
paulson
parents: 33175
diff changeset
  2290
  fixes S :: "real set"
paulson
parents: 33175
diff changeset
  2291
  shows "bounded S ==> Sup(insert x S) = (if S = {} then x else max x (Sup S))" 
paulson
parents: 33175
diff changeset
  2292
by auto (metis Int_absorb Sup_insert_nonempty assms bounded_has_Sup(1) disjoint_iff_not_equal) 
paulson
parents: 33175
diff changeset
  2293
paulson
parents: 33175
diff changeset
  2294
lemma Sup_insert_finite:
paulson
parents: 33175
diff changeset
  2295
  fixes S :: "real set"
paulson
parents: 33175
diff changeset
  2296
  shows "finite S \<Longrightarrow> Sup(insert x S) = (if S = {} then x else max x (Sup S))"
paulson
parents: 33175
diff changeset
  2297
  apply (rule Sup_insert)
paulson
parents: 33175
diff changeset
  2298
  apply (rule finite_imp_bounded)
paulson
parents: 33175
diff changeset
  2299
  by simp
paulson
parents: 33175
diff changeset
  2300
paulson
parents: 33175
diff changeset
  2301
lemma bounded_has_Inf:
paulson
parents: 33175
diff changeset
  2302
  fixes S :: "real set"
paulson
parents: 33175
diff changeset
  2303
  assumes "bounded S"  "S \<noteq> {}"
paulson
parents: 33175
diff changeset
  2304
  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
  2305
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2306
  fix x assume "x\<in>S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2307
  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
  2308
  thus "x \<ge> Inf S" using `x\<in>S`
paulson
parents: 33175
diff changeset
  2309
    by (metis Inf_lower_EX abs_le_D2 minus_le_iff)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2310
next
33270
paulson
parents: 33175
diff changeset
  2311
  show "\<forall>b. (\<forall>x\<in>S. x >= b) \<longrightarrow> Inf S \<ge> b" using assms
paulson
parents: 33175
diff changeset
  2312
    by (metis SupInf.Inf_greatest)
paulson
parents: 33175
diff changeset
  2313
qed
paulson
parents: 33175
diff changeset
  2314
paulson
parents: 33175
diff changeset
  2315
lemma Inf_insert:
paulson
parents: 33175
diff changeset
  2316
  fixes S :: "real set"
paulson
parents: 33175
diff changeset
  2317
  shows "bounded S ==> Inf(insert x S) = (if S = {} then x else min x (Inf S))" 
paulson
parents: 33175
diff changeset
  2318
by auto (metis Int_absorb Inf_insert_nonempty bounded_has_Inf(1) disjoint_iff_not_equal) 
paulson
parents: 33175
diff changeset
  2319
lemma Inf_insert_finite:
paulson
parents: 33175
diff changeset
  2320
  fixes S :: "real set"
paulson
parents: 33175
diff changeset
  2321
  shows "finite S ==> Inf(insert x S) = (if S = {} then x else min x (Inf S))"
paulson
parents: 33175
diff changeset
  2322
  by (rule Inf_insert, rule finite_imp_bounded, simp)
paulson
parents: 33175
diff changeset
  2323
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2324
(* TODO: Move this to RComplete.thy -- would need to include Glb into RComplete *)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2325
lemma real_isGlb_unique: "[| isGlb R S x; isGlb R S y |] ==> x = (y::real)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2326
  apply (frule isGlb_isLb)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2327
  apply (frule_tac x = y in isGlb_isLb)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2328
  apply (blast intro!: order_antisym dest!: isGlb_le_isLb)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2329
  done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2330
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2331
subsection {* Compactness *}
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2332
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2333
subsubsection{* Open-cover compactness *}
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2334
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2335
definition compact :: "'a::topological_space set \<Rightarrow> bool" where
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2336
  compact_eq_heine_borel: -- "This name is used for backwards compatibility"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2337
    "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
  2338
50898
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2339
lemma compactI:
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2340
  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
  2341
  shows "compact s"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2342
  unfolding compact_eq_heine_borel using assms by metis
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2343
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2344
lemma compactE:
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2345
  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
  2346
  obtains C' where "C' \<subseteq> C" and "finite C'" and "s \<subseteq> \<Union>C'"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2347
  using assms unfolding compact_eq_heine_borel by metis
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2348
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2349
subsubsection {* Bolzano-Weierstrass property *}
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2350
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2351
lemma heine_borel_imp_bolzano_weierstrass:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2352
  assumes "compact s" "infinite t"  "t \<subseteq> s"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2353
  shows "\<exists>x \<in> s. x islimpt t"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2354
proof(rule ccontr)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2355
  assume "\<not> (\<exists>x \<in> s. x islimpt t)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2356
  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
  2357
    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
  2358
  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
  2359
    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
  2360
  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
  2361
  { 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
  2362
    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
  2363
    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
  2364
  hence "inj_on f t" unfolding inj_on_def by simp
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2365
  hence "infinite (f ` t)" using assms(2) using finite_imageD by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2366
  moreover
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2367
  { fix x assume "x\<in>t" "f x \<notin> g"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2368
    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
  2369
    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
  2370
    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
  2371
    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
  2372
  hence "f ` t \<subseteq> g" by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2373
  ultimately show False using g(2) using finite_subset by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2374
qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2375
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2376
lemma islimpt_range_imp_convergent_subsequence:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2377
  fixes l :: "'a :: {t1_space, first_countable_topology}"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2378
  assumes l: "l islimpt (range f)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2379
  shows "\<exists>r. subseq r \<and> ((f \<circ> r) ---> l) sequentially"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2380
proof -
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2381
  from first_countable_topology_class.countable_basis_at_decseq[of l] guess A . note A = this
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2382
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2383
  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
  2384
  { fix n i
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2385
    have "\<exists>a. i < a \<and> f a \<in> A (Suc n) - (f ` {.. i} - {l})" (is "\<exists>a. _ \<and> _ \<in> ?A")
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2386
      apply (rule l[THEN islimptE, of "?A"])
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2387
      using A(2) apply fastforce
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2388
      using A(1)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2389
      apply (intro open_Diff finite_imp_closed)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2390
      apply auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2391
      apply (rule_tac x=x in exI)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2392
      apply auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2393
      done
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2394
    then have "\<exists>a. i < a \<and> f a \<in> A (Suc n)" by blast
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2395
    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
  2396
      unfolding s_def by (auto intro: someI2_ex) }
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2397
  note s = this
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2398
  def r \<equiv> "nat_rec (s 0 0) s"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2399
  have "subseq r"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2400
    by (auto simp: r_def s subseq_Suc_iff)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2401
  moreover
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2402
  have "(\<lambda>n. f (r n)) ----> l"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2403
  proof (rule topological_tendstoI)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2404
    fix S assume "open S" "l \<in> S"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2405
    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
  2406
    moreover
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2407
    { 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
  2408
        by (cases i) (simp_all add: r_def s) }
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2409
    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
  2410
    ultimately show "eventually (\<lambda>i. f (r i) \<in> S) sequentially"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2411
      by eventually_elim auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2412
  qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2413
  ultimately show "\<exists>r. subseq r \<and> (f \<circ> r) ----> l"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2414
    by (auto simp: convergent_def comp_def)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2415
qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2416
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2417
lemma finite_range_imp_infinite_repeats:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2418
  fixes f :: "nat \<Rightarrow> 'a"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2419
  assumes "finite (range f)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2420
  shows "\<exists>k. infinite {n. f n = k}"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2421
proof -
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2422
  { fix A :: "'a set" assume "finite A"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2423
    hence "\<And>f. infinite {n. f n \<in> A} \<Longrightarrow> \<exists>k. infinite {n. f n = k}"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2424
    proof (induct)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2425
      case empty thus ?case by simp
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2426
    next
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2427
      case (insert x A)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2428
     show ?case
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2429
      proof (cases "finite {n. f n = x}")
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2430
        case True
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2431
        with `infinite {n. f n \<in> insert x A}`
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2432
        have "infinite {n. f n \<in> A}" by simp
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2433
        thus "\<exists>k. infinite {n. f n = k}" by (rule insert)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2434
      next
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2435
        case False thus "\<exists>k. infinite {n. f n = k}" ..
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2436
      qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2437
    qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2438
  } note H = this
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2439
  from assms show "\<exists>k. infinite {n. f n = k}"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2440
    by (rule H) simp
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2441
qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2442
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2443
lemma sequence_infinite_lemma:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2444
  fixes f :: "nat \<Rightarrow> 'a::t1_space"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2445
  assumes "\<forall>n. f n \<noteq> l" and "(f ---> l) sequentially"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2446
  shows "infinite (range f)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2447
proof
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2448
  assume "finite (range f)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2449
  hence "closed (range f)" by (rule finite_imp_closed)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2450
  hence "open (- range f)" by (rule open_Compl)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2451
  from assms(1) have "l \<in> - range f" by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2452
  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
  2453
    using `open (- range f)` `l \<in> - range f` by (rule topological_tendstoD)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2454
  thus False unfolding eventually_sequentially by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2455
qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2456
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2457
lemma closure_insert:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2458
  fixes x :: "'a::t1_space"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2459
  shows "closure (insert x s) = insert x (closure s)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2460
apply (rule closure_unique)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2461
apply (rule insert_mono [OF closure_subset])
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2462
apply (rule closed_insert [OF closed_closure])
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2463
apply (simp add: closure_minimal)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2464
done
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2465
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2466
lemma islimpt_insert:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2467
  fixes x :: "'a::t1_space"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2468
  shows "x islimpt (insert a s) \<longleftrightarrow> x islimpt s"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2469
proof
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2470
  assume *: "x islimpt (insert a s)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2471
  show "x islimpt s"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2472
  proof (rule islimptI)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2473
    fix t assume t: "x \<in> t" "open t"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2474
    show "\<exists>y\<in>s. y \<in> t \<and> y \<noteq> x"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2475
    proof (cases "x = a")
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2476
      case True
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2477
      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
  2478
        using * t by (rule islimptE)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2479
      with `x = a` show ?thesis by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2480
    next
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2481
      case False
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2482
      with t have t': "x \<in> t - {a}" "open (t - {a})"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2483
        by (simp_all add: open_Diff)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2484
      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
  2485
        using * t' by (rule islimptE)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2486
      thus ?thesis by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2487
    qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2488
  qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2489
next
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2490
  assume "x islimpt s" thus "x islimpt (insert a s)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2491
    by (rule islimpt_subset) auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2492
qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2493
50897
078590669527 generalize lemma islimpt_finite to class t1_space
huffman
parents: 50884
diff changeset
  2494
lemma islimpt_finite:
078590669527 generalize lemma islimpt_finite to class t1_space
huffman
parents: 50884
diff changeset
  2495
  fixes x :: "'a::t1_space"
078590669527 generalize lemma islimpt_finite to class t1_space
huffman
parents: 50884
diff changeset
  2496
  shows "finite s \<Longrightarrow> \<not> x islimpt s"
078590669527 generalize lemma islimpt_finite to class t1_space
huffman
parents: 50884
diff changeset
  2497
by (induct set: finite, simp_all add: islimpt_insert)
078590669527 generalize lemma islimpt_finite to class t1_space
huffman
parents: 50884
diff changeset
  2498
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2499
lemma islimpt_union_finite:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2500
  fixes x :: "'a::t1_space"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2501
  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
  2502
by (simp add: islimpt_Un islimpt_finite)
078590669527 generalize lemma islimpt_finite to class t1_space
huffman
parents: 50884
diff changeset
  2503
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2504
lemma sequence_unique_limpt:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2505
  fixes f :: "nat \<Rightarrow> 'a::t2_space"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2506
  assumes "(f ---> l) sequentially" and "l' islimpt (range f)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2507
  shows "l' = l"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2508
proof (rule ccontr)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2509
  assume "l' \<noteq> l"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2510
  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
  2511
    using hausdorff [OF `l' \<noteq> l`] by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2512
  have "eventually (\<lambda>n. f n \<in> t) sequentially"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2513
    using assms(1) `open t` `l \<in> t` by (rule topological_tendstoD)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2514
  then obtain N where "\<forall>n\<ge>N. f n \<in> t"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2515
    unfolding eventually_sequentially by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2516
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2517
  have "UNIV = {..<N} \<union> {N..}" by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2518
  hence "l' islimpt (f ` ({..<N} \<union> {N..}))" using assms(2) by simp
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2519
  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
  2520
  hence "l' islimpt (f ` {N..})" by (simp add: islimpt_union_finite)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2521
  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
  2522
    using `l' \<in> s` `open s` by (rule islimptE)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2523
  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
  2524
  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
  2525
  with `s \<inter> t = {}` show False by simp
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2526
qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2527
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2528
lemma bolzano_weierstrass_imp_closed:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2529
  fixes s :: "'a::{first_countable_topology, t2_space} set"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2530
  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
  2531
  shows "closed s"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2532
proof-
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2533
  { 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
  2534
    hence "l \<in> s"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2535
    proof(cases "\<forall>n. x n \<noteq> l")
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2536
      case False thus "l\<in>s" using as(1) by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2537
    next
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2538
      case True note cas = this
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2539
      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
  2540
      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
  2541
      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
  2542
    qed  }
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2543
  thus ?thesis unfolding closed_sequential_limits by fast
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2544
qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2545
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2546
lemma compact_imp_closed:
50898
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2547
  fixes s :: "'a::t2_space set"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2548
  assumes "compact s" shows "closed s"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2549
unfolding closed_def
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2550
proof (rule openI)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2551
  fix y assume "y \<in> - s"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2552
  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
  2553
  note `compact s`
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2554
  moreover have "\<forall>u\<in>?C. open u" by simp
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2555
  moreover have "s \<subseteq> \<Union>?C"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2556
  proof
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2557
    fix x assume "x \<in> s"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2558
    with `y \<in> - s` have "x \<noteq> y" by clarsimp
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2559
    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
  2560
      by (rule hausdorff)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2561
    with `x \<in> s` show "x \<in> \<Union>?C"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2562
      unfolding eventually_nhds by auto
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2563
  qed
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2564
  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
  2565
    by (rule compactE)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2566
  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
  2567
  with `finite D` have "eventually (\<lambda>y. y \<notin> \<Union>D) (nhds y)"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2568
    by (simp add: eventually_Ball_finite)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2569
  with `s \<subseteq> \<Union>D` have "eventually (\<lambda>y. y \<notin> s) (nhds y)"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2570
    by (auto elim!: eventually_mono [rotated])
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2571
  thus "\<exists>t. open t \<and> y \<in> t \<and> t \<subseteq> - s"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2572
    by (simp add: eventually_nhds subset_eq)
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2573
qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2574
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2575
text{* In particular, some common special cases. *}
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2576
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2577
lemma compact_empty[simp]:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2578
 "compact {}"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2579
  unfolding compact_eq_heine_borel
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2580
  by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2581
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2582
lemma compact_union [intro]:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2583
  assumes "compact s" "compact t" shows " compact (s \<union> t)"
50898
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2584
proof (rule compactI)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2585
  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
  2586
  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
  2587
    unfolding compact_eq_heine_borel by (auto elim!: allE[of _ f]) metis
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2588
  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
  2589
    unfolding compact_eq_heine_borel by (auto elim!: allE[of _ f]) metis
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2590
  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
  2591
    by (auto intro!: exI[of _ "s' \<union> t'"])
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2592
qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2593
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2594
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
  2595
  by (induct set: finite) auto
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_UN [intro]:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2598
  "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
  2599
  unfolding SUP_def by (rule compact_Union) auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2600
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2601
lemma compact_inter_closed [intro]:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2602
  assumes "compact s" and "closed t"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2603
  shows "compact (s \<inter> t)"
50898
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2604
proof (rule compactI)
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2605
  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
  2606
  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
  2607
  moreover from cover have "s \<subseteq> \<Union>(C \<union> {-t})" by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2608
  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
  2609
    using `compact s` unfolding compact_eq_heine_borel by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2610
  then guess D ..
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2611
  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
  2612
    by (intro exI[of _ "D - {-t}"]) auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2613
qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2614
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2615
lemma closed_inter_compact [intro]:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2616
  assumes "closed s" and "compact t"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2617
  shows "compact (s \<inter> t)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2618
  using compact_inter_closed [of t s] assms
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2619
  by (simp add: Int_commute)
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 [intro]:
50898
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  2622
  fixes s t :: "'a :: t2_space set"
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2623
  assumes "compact s" and "compact t"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2624
  shows "compact (s \<inter> t)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2625
  using assms by (intro compact_inter_closed compact_imp_closed)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2626
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2627
lemma compact_sing [simp]: "compact {a}"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2628
  unfolding compact_eq_heine_borel by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2629
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2630
lemma compact_insert [simp]:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2631
  assumes "compact s" shows "compact (insert x s)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2632
proof -
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2633
  have "compact ({x} \<union> s)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2634
    using compact_sing assms by (rule compact_union)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2635
  thus ?thesis by simp
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2636
qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2637
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2638
lemma finite_imp_compact:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2639
  shows "finite s \<Longrightarrow> compact s"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2640
  by (induct set: finite) simp_all
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2641
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2642
lemma open_delete:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2643
  fixes s :: "'a::t1_space set"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2644
  shows "open s \<Longrightarrow> open (s - {x})"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2645
  by (simp add: open_Diff)
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
text{* Finite intersection property *}
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2648
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2649
lemma inj_setminus: "inj_on uminus (A::'a set set)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2650
  by (auto simp: inj_on_def)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2651
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2652
lemma compact_fip:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2653
  "compact U \<longleftrightarrow>
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2654
    (\<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
  2655
  (is "_ \<longleftrightarrow> ?R")
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2656
proof (safe intro!: compact_eq_heine_borel[THEN iffD2])
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2657
  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
  2658
    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
  2659
  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
  2660
    by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2661
  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
  2662
    unfolding compact_eq_heine_borel by (metis subset_image_iff)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2663
  with fi[THEN spec, of B] show False
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2664
    by (auto dest: finite_imageD intro: inj_setminus)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2665
next
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2666
  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
  2667
  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
  2668
    by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2669
  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
  2670
    by (metis subset_image_iff)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2671
  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
  2672
    by  (auto intro!: exI[of _ B] inj_setminus dest: finite_imageD)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2673
qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2674
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2675
lemma compact_imp_fip:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2676
  "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
  2677
    s \<inter> (\<Inter> f) \<noteq> {}"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2678
  unfolding compact_fip by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2679
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2680
text{*Compactness expressed with filters*}
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2681
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2682
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
  2683
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2684
lemma eventually_filter_from_subbase:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2685
  "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
  2686
    (is "_ \<longleftrightarrow> ?R P")
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2687
  unfolding filter_from_subbase_def
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2688
proof (rule eventually_Abs_filter is_filter.intro)+
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2689
  show "?R (\<lambda>x. True)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2690
    by (rule exI[of _ "{}"]) (simp add: le_fun_def)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2691
next
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2692
  fix P Q assume "?R P" then guess X ..
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2693
  moreover assume "?R Q" then guess Y ..
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2694
  ultimately show "?R (\<lambda>x. P x \<and> Q x)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2695
    by (intro exI[of _ "X \<union> Y"]) auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2696
next
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2697
  fix P Q
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2698
  assume "?R P" then guess X ..
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2699
  moreover assume "\<forall>x. P x \<longrightarrow> Q x"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2700
  ultimately show "?R Q"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2701
    by (intro exI[of _ X]) auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2702
qed
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_subbaseI: "P \<in> B \<Longrightarrow> eventually P (filter_from_subbase B)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2705
  by (subst eventually_filter_from_subbase) (auto intro!: exI[of _ "{P}"])
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2706
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2707
lemma filter_from_subbase_not_bot:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2708
  "\<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
  2709
  unfolding trivial_limit_def eventually_filter_from_subbase by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2710
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2711
lemma closure_iff_nhds_not_empty:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2712
  "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
  2713
proof safe
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2714
  assume x: "x \<in> closure X"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2715
  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
  2716
  then have "x \<notin> closure (-S)" 
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2717
    by (auto simp: closure_complement subset_eq[symmetric] intro: interiorI)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2718
  with x have "x \<in> closure X - closure (-S)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2719
    by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2720
  also have "\<dots> \<subseteq> closure (X \<inter> S)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2721
    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
  2722
  finally have "X \<inter> S \<noteq> {}" by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2723
  then show False using `X \<inter> A = {}` `S \<subseteq> A` by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2724
next
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2725
  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
  2726
  from this[THEN spec, of "- X", THEN spec, of "- closure X"]
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2727
  show "x \<in> closure X"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2728
    by (simp add: closure_subset open_Compl)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2729
qed
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 compact_filter:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2732
  "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
  2733
proof (intro allI iffI impI compact_fip[THEN iffD2] notI)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2734
  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
  2735
  from F have "U \<noteq> {}"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2736
    by (auto simp: eventually_False)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2737
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2738
  def Z \<equiv> "closure ` {A. eventually (\<lambda>x. x \<in> A) F}"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2739
  then have "\<forall>z\<in>Z. closed z"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2740
    by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2741
  moreover 
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2742
  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
  2743
    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
  2744
  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
  2745
  proof (intro allI impI)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2746
    fix B assume "finite B" "B \<subseteq> Z"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2747
    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
  2748
      by (auto intro!: eventually_Ball_finite)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2749
    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
  2750
      by eventually_elim auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2751
    with F show "U \<inter> \<Inter>B \<noteq> {}"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2752
      by (intro notI) (simp add: eventually_False)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2753
  qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2754
  ultimately have "U \<inter> \<Inter>Z \<noteq> {}"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2755
    using `compact U` unfolding compact_fip by blast
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2756
  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
  2757
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2758
  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
  2759
    unfolding eventually_inf eventually_nhds
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2760
  proof safe
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2761
    fix P Q R S
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2762
    assume "eventually R F" "open S" "x \<in> S"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2763
    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
  2764
    have "S \<inter> {x. R x} \<noteq> {}" by (auto simp: Z_def)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2765
    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
  2766
    ultimately show False by (auto simp: set_eq_iff)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2767
  qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2768
  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
  2769
    by (metis eventually_bot)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2770
next
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2771
  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
  2772
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2773
  def P' \<equiv> "(\<lambda>a (x::'a). x \<in> a)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2774
  then have inj_P': "\<And>A. inj_on P' A"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2775
    by (auto intro!: inj_onI simp: fun_eq_iff)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2776
  def F \<equiv> "filter_from_subbase (P' ` insert U A)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2777
  have "F \<noteq> bot"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2778
    unfolding F_def
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2779
  proof (safe intro!: filter_from_subbase_not_bot)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2780
    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
  2781
    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
  2782
      unfolding subset_image_iff by (auto intro: inj_P' finite_imageD)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2783
    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
  2784
    with B show False by (auto simp: P'_def fun_eq_iff)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2785
  qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2786
  moreover have "eventually (\<lambda>x. x \<in> U) F"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2787
    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
  2788
  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
  2789
  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
  2790
    by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2791
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2792
  { fix V assume "V \<in> A"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2793
    then have V: "eventually (\<lambda>x. x \<in> V) F"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2794
      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
  2795
    have "x \<in> closure V"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2796
      unfolding closure_iff_nhds_not_empty
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2797
    proof (intro impI allI)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2798
      fix S A assume "open S" "x \<in> S" "S \<subseteq> A"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2799
      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
  2800
      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
  2801
        by (auto simp: eventually_inf)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2802
      with x show "V \<inter> A \<noteq> {}"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2803
        by (auto simp del: Int_iff simp add: trivial_limit_def)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2804
    qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2805
    then have "x \<in> V"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2806
      using `V \<in> A` A(1) by simp }
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2807
  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
  2808
  with `U \<inter> \<Inter>A = {}` show False by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2809
qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2810
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2811
lemma countable_compact:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2812
  fixes U :: "'a :: second_countable_topology set"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2813
  shows "compact U \<longleftrightarrow>
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2814
    (\<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))"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2815
proof (safe intro!: compact_eq_heine_borel[THEN iffD2])
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2816
  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
  2817
  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)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2818
  def B \<equiv> "SOME B::'a set set. countable B \<and> topological_basis B"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2819
  then have B: "countable B" "topological_basis B"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2820
    by (auto simp: countable_basis is_basis)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2821
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2822
  moreover def C \<equiv> "{b\<in>B. \<exists>a\<in>A. b \<subseteq> a}"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2823
  ultimately have "countable C" "\<forall>a\<in>C. open a"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2824
    unfolding C_def by (auto simp: topological_basis_open)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2825
  moreover
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2826
  have "\<Union>A \<subseteq> \<Union>C"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2827
  proof safe
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2828
    fix x a assume "x \<in> a" "a \<in> A"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2829
    with topological_basisE[of B a x] B A
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2830
    obtain b where "x \<in> b" "b \<in> B" "b \<subseteq> a" by metis
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2831
    with `a \<in> A` show "x \<in> \<Union>C" unfolding C_def by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2832
  qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2833
  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
  2834
  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
  2835
    using * by metis
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2836
  moreover then have "\<forall>t\<in>T. \<exists>a\<in>A. t \<subseteq> a"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2837
    by (auto simp: C_def)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2838
  then guess f unfolding bchoice_iff Bex_def ..
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2839
  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
  2840
    unfolding C_def by (intro exI[of _ "f`T"]) fastforce
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2841
qed (auto simp: compact_eq_heine_borel)
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  2842
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  2843
subsubsection{* Sequential compactness *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2844
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2845
definition
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2846
  seq_compact :: "'a::topological_space set \<Rightarrow> bool" where
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2847
  "seq_compact S \<longleftrightarrow>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2848
   (\<forall>f. (\<forall>n. f n \<in> S) \<longrightarrow>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2849
       (\<exists>l\<in>S. \<exists>r. subseq r \<and> ((f o r) ---> l) sequentially))"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2850
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2851
lemma seq_compact_imp_compact:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2852
  fixes U :: "'a :: second_countable_topology set"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2853
  assumes "seq_compact U"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2854
  shows "compact U"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2855
  unfolding countable_compact
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2856
proof safe
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2857
  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
  2858
  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
  2859
    using `seq_compact U` by (fastforce simp: seq_compact_def subset_eq)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2860
  show "\<exists>T\<subseteq>A. finite T \<and> U \<subseteq> \<Union>T"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2861
  proof cases
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2862
    assume "finite A" with A show ?thesis by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2863
  next
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2864
    assume "infinite A"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2865
    then have "A \<noteq> {}" by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2866
    show ?thesis
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2867
    proof (rule ccontr)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2868
      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
  2869
      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
  2870
      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
  2871
      def X \<equiv> "\<lambda>n. X' (from_nat_into A ` {.. n})"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2872
      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
  2873
        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
  2874
      then have "range X \<subseteq> U" by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2875
      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
  2876
      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
  2877
      obtain n where "x \<in> from_nat_into A n" by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2878
      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
  2879
      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
  2880
        unfolding tendsto_def by (auto simp: comp_def)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2881
      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
  2882
        by (auto simp: eventually_sequentially)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2883
      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
  2884
        by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2885
      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
  2886
        by (auto intro!: exI[of _ "max n N"])
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2887
      ultimately show False
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2888
        by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2889
    qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2890
  qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2891
qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2892
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2893
lemma compact_imp_seq_compact:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2894
  fixes U :: "'a :: first_countable_topology set"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2895
  assumes "compact U" shows "seq_compact U"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2896
  unfolding seq_compact_def
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2897
proof safe
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2898
  fix X :: "nat \<Rightarrow> 'a" assume "\<forall>n. X n \<in> U"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2899
  then have "eventually (\<lambda>x. x \<in> U) (filtermap X sequentially)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2900
    by (auto simp: eventually_filtermap)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2901
  moreover have "filtermap X sequentially \<noteq> bot"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2902
    by (simp add: trivial_limit_def eventually_filtermap)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2903
  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
  2904
    using `compact U` by (auto simp: compact_filter)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2905
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2906
  from countable_basis_at_decseq[of x] guess A . note A = this
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2907
  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
  2908
  { fix n i
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2909
    have "\<exists>a. i < a \<and> X a \<in> A (Suc n)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2910
    proof (rule ccontr)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2911
      assume "\<not> (\<exists>a>i. X a \<in> A (Suc n))"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2912
      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
  2913
      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
  2914
        by (auto simp: eventually_filtermap eventually_sequentially)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2915
      moreover have "eventually (\<lambda>x. x \<in> A (Suc n)) (nhds x)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2916
        using A(1,2)[of "Suc n"] by (auto simp: eventually_nhds)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2917
      ultimately have "eventually (\<lambda>x. False) ?F"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2918
        by (auto simp add: eventually_inf)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2919
      with x show False
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2920
        by (simp add: eventually_False)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2921
    qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2922
    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
  2923
      unfolding s_def by (auto intro: someI2_ex) }
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2924
  note s = this
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2925
  def r \<equiv> "nat_rec (s 0 0) s"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2926
  have "subseq r"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2927
    by (auto simp: r_def s subseq_Suc_iff)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2928
  moreover
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2929
  have "(\<lambda>n. X (r n)) ----> x"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2930
  proof (rule topological_tendstoI)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2931
    fix S assume "open S" "x \<in> S"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2932
    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
  2933
    moreover
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2934
    { 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
  2935
        by (cases i) (simp_all add: r_def s) }
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2936
    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
  2937
    ultimately show "eventually (\<lambda>i. X (r i) \<in> S) sequentially"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2938
      by eventually_elim auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2939
  qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2940
  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
  2941
    using `x \<in> U` by (auto simp: convergent_def comp_def)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2942
qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2943
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2944
lemma seq_compact_eq_compact:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2945
  fixes U :: "'a :: second_countable_topology set"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2946
  shows "seq_compact U \<longleftrightarrow> compact U"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2947
  using compact_imp_seq_compact seq_compact_imp_compact by blast
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2948
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2949
lemma seq_compactI:
44075
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  2950
  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
  2951
  shows "seq_compact S"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2952
  unfolding seq_compact_def using assms by fast
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2953
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2954
lemma seq_compactE:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2955
  assumes "seq_compact S" "\<forall>n. f n \<in> S"
44075
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  2956
  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
  2957
  using assms unfolding seq_compact_def by fast
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2958
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2959
lemma bolzano_weierstrass_imp_seq_compact:
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2960
  fixes s :: "'a::{t1_space, first_countable_topology} set"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2961
  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
  2962
  shows "seq_compact s"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2963
proof -
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2964
  { fix f :: "nat \<Rightarrow> 'a" assume f: "\<forall>n. f n \<in> s"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2965
    have "\<exists>l\<in>s. \<exists>r. subseq r \<and> ((f \<circ> r) ---> l) sequentially"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2966
    proof (cases "finite (range f)")
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2967
      case True
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2968
      hence "\<exists>l. infinite {n. f n = l}"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2969
        by (rule finite_range_imp_infinite_repeats)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2970
      then obtain l where "infinite {n. f n = l}" ..
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2971
      hence "\<exists>r. subseq r \<and> (\<forall>n. r n \<in> {n. f n = l})"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2972
        by (rule infinite_enumerate)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2973
      then obtain r where "subseq r" and fr: "\<forall>n. f (r n) = l" by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2974
      hence "subseq r \<and> ((f \<circ> r) ---> l) sequentially"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2975
        unfolding o_def by (simp add: fr tendsto_const)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2976
      hence "\<exists>r. subseq r \<and> ((f \<circ> r) ---> l) sequentially"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2977
        by - (rule exI)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2978
      from f have "\<forall>n. f (r n) \<in> s" by simp
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2979
      hence "l \<in> s" by (simp add: fr)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2980
      thus "\<exists>l\<in>s. \<exists>r. subseq r \<and> ((f \<circ> r) ---> l) sequentially"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2981
        by (rule rev_bexI) fact
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2982
    next
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2983
      case False
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2984
      with f assms have "\<exists>x\<in>s. x islimpt (range f)" by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2985
      then obtain l where "l \<in> s" "l islimpt (range f)" ..
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2986
      have "\<exists>r. subseq r \<and> ((f \<circ> r) ---> l) sequentially"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2987
        using `l islimpt (range f)`
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2988
        by (rule islimpt_range_imp_convergent_subsequence)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2989
      with `l \<in> s` show "\<exists>l\<in>s. \<exists>r. subseq r \<and> ((f \<circ> r) ---> l) sequentially" ..
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2990
    qed
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2991
  }
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2992
  thus ?thesis unfolding seq_compact_def by auto
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  2993
qed
44075
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  2994
50940
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  2995
subsubsection{* Total boundedness *}
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  2996
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  2997
lemma cauchy_def:
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  2998
  "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
  2999
unfolding Cauchy_def by blast
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3000
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3001
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
  3002
  "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
  3003
declare helper_1.simps[simp del]
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3004
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3005
lemma seq_compact_imp_totally_bounded:
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3006
  assumes "seq_compact s"
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3007
  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
  3008
proof(rule, rule, rule ccontr)
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3009
  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
  3010
  def x \<equiv> "helper_1 s e"
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3011
  { fix n
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3012
    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
  3013
    proof(induct_tac rule:nat_less_induct)
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3014
      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
  3015
      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
  3016
      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
  3017
      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
  3018
      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
  3019
        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
  3020
      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
  3021
    qed }
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3022
  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
  3023
  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
  3024
  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
  3025
  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
  3026
  show False
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3027
    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
  3028
    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
  3029
    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
  3030
qed
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3031
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3032
subsubsection{* Heine-Borel theorem *}
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3033
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3034
text {* Following Burkill \& Burkill vol. 2. *}
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3035
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3036
lemma heine_borel_lemma: fixes s::"'a::metric_space set"
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3037
  assumes "seq_compact s"  "s \<subseteq> (\<Union> t)"  "\<forall>b \<in> t. open b"
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3038
  shows "\<exists>e>0. \<forall>x \<in> s. \<exists>b \<in> t. ball x e \<subseteq> b"
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3039
proof(rule ccontr)
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3040
  assume "\<not> (\<exists>e>0. \<forall>x\<in>s. \<exists>b\<in>t. ball x e \<subseteq> b)"
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3041
  hence cont:"\<forall>e>0. \<exists>x\<in>s. \<forall>xa\<in>t. \<not> (ball x e \<subseteq> xa)" by auto
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3042
  { fix n::nat
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3043
    have "1 / real (n + 1) > 0" by auto
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3044
    hence "\<exists>x. x\<in>s \<and> (\<forall>xa\<in>t. \<not> (ball x (inverse (real (n+1))) \<subseteq> xa))" using cont unfolding Bex_def by auto }
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3045
  hence "\<forall>n::nat. \<exists>x. x \<in> s \<and> (\<forall>xa\<in>t. \<not> ball x (inverse (real (n + 1))) \<subseteq> xa)" by auto
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3046
  then obtain f where f:"\<forall>n::nat. f n \<in> s \<and> (\<forall>xa\<in>t. \<not> ball (f n) (inverse (real (n + 1))) \<subseteq> xa)"
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3047
    using choice[of "\<lambda>n::nat. \<lambda>x. x\<in>s \<and> (\<forall>xa\<in>t. \<not> ball x (inverse (real (n + 1))) \<subseteq> xa)"] by auto
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3048
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3049
  then obtain l r where l:"l\<in>s" and r:"subseq r" and lr:"((f \<circ> r) ---> l) sequentially"
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3050
    using assms(1)[unfolded seq_compact_def, THEN spec[where x=f]] by auto
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3051
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3052
  obtain b where "l\<in>b" "b\<in>t" using assms(2) and l by auto
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3053
  then obtain e where "e>0" and e:"\<forall>z. dist z l < e \<longrightarrow> z\<in>b"
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3054
    using assms(3)[THEN bspec[where x=b]] unfolding open_dist by auto
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3055
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3056
  then obtain N1 where N1:"\<forall>n\<ge>N1. dist ((f \<circ> r) n) l < e / 2"
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3057
    using lr[unfolded LIMSEQ_def, THEN spec[where x="e/2"]] by auto
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3058
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3059
  obtain N2::nat where N2:"N2>0" "inverse (real N2) < e /2" using real_arch_inv[of "e/2"] and `e>0` by auto
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3060
  have N2':"inverse (real (r (N1 + N2) +1 )) < e/2"
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3061
    apply(rule order_less_trans) apply(rule less_imp_inverse_less) using N2
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3062
    using seq_suble[OF r, of "N1 + N2"] by auto
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3063
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3064
  def x \<equiv> "(f (r (N1 + N2)))"
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3065
  have x:"\<not> ball x (inverse (real (r (N1 + N2) + 1))) \<subseteq> b" unfolding x_def
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3066
    using f[THEN spec[where x="r (N1 + N2)"]] using `b\<in>t` by auto
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3067
  have "\<exists>y\<in>ball x (inverse (real (r (N1 + N2) + 1))). y\<notin>b" apply(rule ccontr) using x by auto
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3068
  then obtain y where y:"y \<in> ball x (inverse (real (r (N1 + N2) + 1)))" "y \<notin> b" by auto
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3069
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3070
  have "dist x l < e/2" using N1 unfolding x_def o_def by auto
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3071
  hence "dist y l < e" using y N2' using dist_triangle[of y l x]by (auto simp add:dist_commute)
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3072
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3073
  thus False using e and `y\<notin>b` by auto
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3074
qed
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3075
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3076
lemma seq_compact_imp_heine_borel:
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3077
  fixes s :: "'a :: metric_space set"
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3078
  shows "seq_compact s \<Longrightarrow> compact s"
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3079
  unfolding compact_eq_heine_borel
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3080
proof clarify
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3081
  fix f assume "seq_compact s" " \<forall>t\<in>f. open t" "s \<subseteq> \<Union>f"
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3082
  then obtain e::real where "e>0" and "\<forall>x\<in>s. \<exists>b\<in>f. ball x e \<subseteq> b" using heine_borel_lemma[of s f] by auto
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3083
  hence "\<forall>x\<in>s. \<exists>b. b\<in>f \<and> ball x e \<subseteq> b" by auto
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3084
  hence "\<exists>bb. \<forall>x\<in>s. bb x \<in>f \<and> ball x e \<subseteq> bb x" using bchoice[of s "\<lambda>x b. b\<in>f \<and> ball x e \<subseteq> b"] by auto
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3085
  then obtain  bb where bb:"\<forall>x\<in>s. (bb x) \<in> f \<and> ball x e \<subseteq> (bb x)" by blast
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3086
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3087
  from `seq_compact s` have  "\<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
  3088
    using seq_compact_imp_totally_bounded[of s] `e>0` by auto
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3089
  then obtain k where k:"finite k" "k \<subseteq> s" "s \<subseteq> \<Union>(\<lambda>x. ball x e) ` k" by auto
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
  have "finite (bb ` k)" using k(1) by auto
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3092
  moreover
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3093
  { fix x assume "x\<in>s"
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3094
    hence "x\<in>\<Union>(\<lambda>x. ball x e) ` k" using k(3)  unfolding subset_eq by auto
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3095
    hence "\<exists>X\<in>bb ` k. x \<in> X" using bb k(2) by blast
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3096
    hence "x \<in> \<Union>(bb ` k)" using  Union_iff[of x "bb ` k"] by auto
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3097
  }
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3098
  ultimately show "\<exists>f'\<subseteq>f. finite f' \<and> s \<subseteq> \<Union>f'" using bb k(2) by (rule_tac x="bb ` k" in exI) auto
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3099
qed
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3100
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3101
lemma compact_eq_seq_compact_metric:
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3102
  "compact (s :: 'a::metric_space set) \<longleftrightarrow> seq_compact s"
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3103
  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
  3104
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3105
lemma compact_def:
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3106
  "compact (S :: 'a::metric_space set) \<longleftrightarrow>
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3107
   (\<forall>f. (\<forall>n. f n \<in> S) \<longrightarrow>
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3108
       (\<exists>l\<in>S. \<exists>r. subseq r \<and> ((f o r) ---> l) sequentially)) "
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3109
  unfolding compact_eq_seq_compact_metric seq_compact_def by auto
a7c273a83d27 group compactness-eq-seq-compactness lemmas together
hoelzl
parents: 50939
diff changeset
  3110
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3111
text {*
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3112
  A metric space (or topological vector space) is said to have the
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3113
  Heine-Borel property if every closed and bounded subset is compact.
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3114
*}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3115
44207
ea99698c2070 locale-ize some definitions, so perfect_space and heine_borel can inherit from the proper superclasses
huffman
parents: 44170
diff changeset
  3116
class heine_borel = metric_space +
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3117
  assumes bounded_imp_convergent_subsequence:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3118
    "bounded s \<Longrightarrow> \<forall>n. f n \<in> s
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3119
      \<Longrightarrow> \<exists>l r. subseq r \<and> ((f \<circ> r) ---> l) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3120
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  3121
lemma bounded_closed_imp_seq_compact:
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3122
  fixes s::"'a::heine_borel set"
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  3123
  assumes "bounded s" and "closed s" shows "seq_compact s"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  3124
proof (unfold seq_compact_def, clarify)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3125
  fix f :: "nat \<Rightarrow> 'a" assume f: "\<forall>n. f n \<in> s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3126
  obtain l r where r: "subseq r" and l: "((f \<circ> r) ---> l) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3127
    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
  3128
  from f have fr: "\<forall>n. (f \<circ> r) n \<in> s" by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3129
  have "l \<in> s" using `closed s` fr l
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3130
    unfolding closed_sequential_limits by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3131
  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
  3132
    using `l \<in> s` r l by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3133
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3134
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3135
lemma lim_subseq:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3136
  "subseq r \<Longrightarrow> (s ---> l) sequentially \<Longrightarrow> ((s o r) ---> l) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3137
unfolding tendsto_def eventually_sequentially o_def
50937
d249ef928ae1 removed subseq_bigger (replaced by seq_suble)
hoelzl
parents: 50936
diff changeset
  3138
by (metis seq_suble le_trans)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3139
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3140
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
  3141
  unfolding Ex1_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3142
  apply (rule_tac x="nat_rec e f" in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3143
  apply (rule conjI)+
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3144
apply (rule def_nat_rec_0, simp)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3145
apply (rule allI, rule def_nat_rec_Suc, simp)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3146
apply (rule allI, rule impI, rule ext)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3147
apply (erule conjE)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3148
apply (induct_tac x)
36362
06475a1547cb fix lots of looping simp calls and other warnings
huffman
parents: 36360
diff changeset
  3149
apply simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3150
apply (erule_tac x="n" in allE)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3151
apply (simp)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3152
done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3153
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3154
lemma convergent_bounded_increasing: fixes s ::"nat\<Rightarrow>real"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3155
  assumes "incseq s" and "\<forall>n. abs(s n) \<le> b"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3156
  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
  3157
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3158
  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
  3159
  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
  3160
  { 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
  3161
    { fix n::nat
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3162
      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
  3163
      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
  3164
      with n have "s N \<le> t - e" using `e>0` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3165
      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
  3166
    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
  3167
    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
  3168
  thus ?thesis by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3169
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3170
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3171
lemma convergent_bounded_monotone: fixes s::"nat \<Rightarrow> real"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3172
  assumes "\<forall>n. abs(s n) \<le> b" and "monoseq s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3173
  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
  3174
  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
  3175
  unfolding monoseq_def incseq_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3176
  apply auto unfolding minus_add_distrib[THEN sym, unfolded diff_minus[THEN sym]]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3177
  unfolding abs_minus_cancel by(rule_tac x="-l" in exI)auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3178
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  3179
(* 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
  3180
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
  3181
  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
  3182
  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
  3183
proof-
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  3184
  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
  3185
  { 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
  3186
    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
  3187
      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
  3188
      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
  3189
  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
  3190
  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
  3191
    unfolding monoseq_def by auto
44907
93943da0a010 remove redundant lemma Lim_sequentially in favor of lemma LIMSEQ_def
huffman
parents: 44905
diff changeset
  3192
  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
  3193
    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
  3194
qed
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  3195
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3196
lemma compact_real_lemma:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3197
  assumes "\<forall>n::nat. abs(s n) \<le> b"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3198
  shows "\<exists>(l::real) r. subseq r \<and> ((s \<circ> r) ---> l) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3199
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3200
  obtain r where r:"subseq r" "monoseq (\<lambda>n. s (r n))"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3201
    using seq_monosub[of s] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3202
  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
  3203
    unfolding tendsto_iff dist_norm eventually_sequentially by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3204
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3205
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3206
instance real :: heine_borel
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3207
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3208
  fix s :: "real set" and f :: "nat \<Rightarrow> real"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3209
  assume s: "bounded s" and f: "\<forall>n. f n \<in> s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3210
  then obtain b where b: "\<forall>n. abs (f n) \<le> b"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3211
    unfolding bounded_iff by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3212
  obtain l :: real and r :: "nat \<Rightarrow> nat" where
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3213
    r: "subseq r" and l: "((f \<circ> r) ---> l) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3214
    using compact_real_lemma [OF b] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3215
  thus "\<exists>l r. subseq r \<and> ((f \<circ> r) ---> l) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3216
    by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3217
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3218
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3219
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
  3220
  fixes f :: "nat \<Rightarrow> 'a::euclidean_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3221
  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
  3222
  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
  3223
        (\<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
  3224
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
  3225
  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
  3226
  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
  3227
  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
  3228
      (\<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
  3229
  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
  3230
  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
  3231
    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
  3232
      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
  3233
    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
  3234
      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
  3235
      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
  3236
    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
  3237
    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
  3238
      using bounded_imp_convergent_subsequence[OF s' f'] unfolding o_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3239
    def r \<equiv> "r1 \<circ> r2" have r:"subseq r"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3240
      using r1 and r2 unfolding r_def o_def subseq_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3241
    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
  3242
    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
  3243
    { 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
  3244
      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
  3245
      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
  3246
      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
  3247
        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
  3248
      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
  3249
        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
  3250
        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
  3251
    }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3252
    ultimately show ?case by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3253
  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
  3254
qed
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  3255
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  3256
instance euclidean_space \<subseteq> heine_borel
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3257
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
  3258
  fix s :: "'a set" and f :: "nat \<Rightarrow> 'a"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3259
  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
  3260
  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
  3261
    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
  3262
    using compact_lemma [OF s f] by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3263
  { 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
  3264
    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
  3265
    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
  3266
      by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3267
    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
  3268
    { 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
  3269
      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
  3270
        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
  3271
      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
  3272
        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
  3273
      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
  3274
        by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3275
    }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3276
    ultimately have "eventually (\<lambda>n. dist (f (r n)) l < e) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3277
      by (rule eventually_elim1)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3278
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3279
  hence *:"((f \<circ> r) ---> l) sequentially" unfolding o_def tendsto_iff by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3280
  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
  3281
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3282
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3283
lemma bounded_fst: "bounded s \<Longrightarrow> bounded (fst ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3284
unfolding bounded_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3285
apply clarify
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3286
apply (rule_tac x="a" in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3287
apply (rule_tac x="e" in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3288
apply clarsimp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3289
apply (drule (1) bspec)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3290
apply (simp add: dist_Pair_Pair)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3291
apply (erule order_trans [OF real_sqrt_sum_squares_ge1])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3292
done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3293
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3294
lemma bounded_snd: "bounded s \<Longrightarrow> bounded (snd ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3295
unfolding bounded_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3296
apply clarify
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3297
apply (rule_tac x="b" in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3298
apply (rule_tac x="e" in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3299
apply clarsimp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3300
apply (drule (1) bspec)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3301
apply (simp add: dist_Pair_Pair)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3302
apply (erule order_trans [OF real_sqrt_sum_squares_ge2])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3303
done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3304
37678
0040bafffdef "prod" and "sum" replace "*" and "+" respectively
haftmann
parents: 37649
diff changeset
  3305
instance prod :: (heine_borel, heine_borel) heine_borel
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3306
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3307
  fix s :: "('a * 'b) set" and f :: "nat \<Rightarrow> 'a * 'b"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3308
  assume s: "bounded s" and f: "\<forall>n. f n \<in> s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3309
  from s have s1: "bounded (fst ` s)" by (rule bounded_fst)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3310
  from f have f1: "\<forall>n. fst (f n) \<in> fst ` s" by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3311
  obtain l1 r1 where r1: "subseq r1"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3312
    and l1: "((\<lambda>n. fst (f (r1 n))) ---> l1) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3313
    using bounded_imp_convergent_subsequence [OF s1 f1]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3314
    unfolding o_def by fast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3315
  from s have s2: "bounded (snd ` s)" by (rule bounded_snd)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3316
  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
  3317
  obtain l2 r2 where r2: "subseq r2"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3318
    and l2: "((\<lambda>n. snd (f (r1 (r2 n)))) ---> l2) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3319
    using bounded_imp_convergent_subsequence [OF s2 f2]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3320
    unfolding o_def by fast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3321
  have l1': "((\<lambda>n. fst (f (r1 (r2 n)))) ---> l1) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3322
    using lim_subseq [OF r2 l1] unfolding o_def .
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3323
  have l: "((f \<circ> (r1 \<circ> r2)) ---> (l1, l2)) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3324
    using tendsto_Pair [OF l1' l2] unfolding o_def by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3325
  have r: "subseq (r1 \<circ> r2)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3326
    using r1 r2 unfolding subseq_def by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3327
  show "\<exists>l r. subseq r \<and> ((f \<circ> r) ---> l) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3328
    using l r by fast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3329
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3330
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  3331
subsubsection{* Completeness *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3332
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3333
definition
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3334
  complete :: "'a::metric_space set \<Rightarrow> bool" where
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3335
  "complete s \<longleftrightarrow> (\<forall>f. (\<forall>n. f n \<in> s) \<and> Cauchy f
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3336
                      --> (\<exists>l \<in> s. (f ---> l) sequentially))"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3337
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3338
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
  3339
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3340
  { assume ?rhs
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3341
    { fix e::real
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3342
      assume "e>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3343
      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
  3344
        by (erule_tac x="e/2" in allE) auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3345
      { fix n m
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3346
        assume nm:"N \<le> m \<and> N \<le> n"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3347
        hence "dist (s m) (s n) < e" using N
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3348
          using dist_triangle_half_l[of "s m" "s N" "e" "s n"]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3349
          by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3350
      }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3351
      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
  3352
        by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3353
    }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3354
    hence ?lhs
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3355
      unfolding cauchy_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3356
      by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3357
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3358
  thus ?thesis
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3359
    unfolding cauchy_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3360
    using dist_triangle_half_l
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3361
    by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3362
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3363
34104
22758f95e624 re-state lemmas using 'range'
huffman
parents: 33758
diff changeset
  3364
lemma cauchy_imp_bounded: assumes "Cauchy s" shows "bounded (range s)"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3365
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3366
  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
  3367
  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
  3368
  moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3369
  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
  3370
  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
  3371
    unfolding bounded_any_center [where a="s N"] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3372
  ultimately show "?thesis"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3373
    unfolding bounded_any_center [where a="s N"]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3374
    apply(rule_tac x="max a 1" in exI) apply auto
34104
22758f95e624 re-state lemmas using 'range'
huffman
parents: 33758
diff changeset
  3375
    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
  3376
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3377
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  3378
lemma seq_compact_imp_complete: assumes "seq_compact s" shows "complete s"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3379
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3380
  { 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
  3381
    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
  3382
50937
d249ef928ae1 removed subseq_bigger (replaced by seq_suble)
hoelzl
parents: 50936
diff changeset
  3383
    note lr' = seq_suble [OF lr(2)]
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3384
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3385
    { fix e::real assume "e>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3386
      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
  3387
      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
  3388
      { fix n::nat assume n:"n \<ge> max N M"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3389
        have "dist ((f \<circ> r) n) l < e/2" using n M by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3390
        moreover have "r n \<ge> N" using lr'[of n] n by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3391
        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
  3392
        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
  3393
      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
  3394
    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
  3395
  thus ?thesis unfolding complete_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3396
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3397
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3398
instance heine_borel < complete_space
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3399
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3400
  fix f :: "nat \<Rightarrow> 'a" assume "Cauchy f"
34104
22758f95e624 re-state lemmas using 'range'
huffman
parents: 33758
diff changeset
  3401
  hence "bounded (range f)"
22758f95e624 re-state lemmas using 'range'
huffman
parents: 33758
diff changeset
  3402
    by (rule cauchy_imp_bounded)
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  3403
  hence "seq_compact (closure (range f))"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  3404
    using bounded_closed_imp_seq_compact [of "closure (range f)"] by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3405
  hence "complete (closure (range f))"
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  3406
    by (rule seq_compact_imp_complete)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3407
  moreover have "\<forall>n. f n \<in> closure (range f)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3408
    using closure_subset [of "range f"] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3409
  ultimately have "\<exists>l\<in>closure (range f). (f ---> l) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3410
    using `Cauchy f` unfolding complete_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3411
  then show "convergent f"
36660
1cc4ab4b7ff7 make (X ----> L) an abbreviation for (X ---> L) sequentially
huffman
parents: 36659
diff changeset
  3412
    unfolding convergent_def by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3413
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3414
44632
076a45f65e12 simplify/generalize some proofs
huffman
parents: 44628
diff changeset
  3415
instance euclidean_space \<subseteq> banach ..
076a45f65e12 simplify/generalize some proofs
huffman
parents: 44628
diff changeset
  3416
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3417
lemma complete_univ: "complete (UNIV :: 'a::complete_space set)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3418
proof(simp add: complete_def, rule, rule)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3419
  fix f :: "nat \<Rightarrow> 'a" assume "Cauchy f"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3420
  hence "convergent f" by (rule Cauchy_convergent)
36660
1cc4ab4b7ff7 make (X ----> L) an abbreviation for (X ---> L) sequentially
huffman
parents: 36659
diff changeset
  3421
  thus "\<exists>l. f ----> l" unfolding convergent_def .  
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3422
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3423
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3424
lemma complete_imp_closed: assumes "complete s" shows "closed s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3425
proof -
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3426
  { fix x assume "x islimpt s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3427
    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
  3428
      unfolding islimpt_sequential by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3429
    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
  3430
      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
  3431
    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
  3432
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3433
  thus "closed s" unfolding closed_limpt by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3434
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3435
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3436
lemma complete_eq_closed:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3437
  fixes s :: "'a::complete_space set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3438
  shows "complete s \<longleftrightarrow> closed s" (is "?lhs = ?rhs")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3439
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3440
  assume ?lhs thus ?rhs by (rule complete_imp_closed)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3441
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3442
  assume ?rhs
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3443
  { fix f assume as:"\<forall>n::nat. f n \<in> s" "Cauchy f"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3444
    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
  3445
    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
  3446
  thus ?lhs unfolding complete_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3447
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3448
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3449
lemma convergent_eq_cauchy:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3450
  fixes s :: "nat \<Rightarrow> 'a::complete_space"
44632
076a45f65e12 simplify/generalize some proofs
huffman
parents: 44628
diff changeset
  3451
  shows "(\<exists>l. (s ---> l) sequentially) \<longleftrightarrow> Cauchy s"
076a45f65e12 simplify/generalize some proofs
huffman
parents: 44628
diff changeset
  3452
  unfolding Cauchy_convergent_iff convergent_def ..
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3453
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3454
lemma convergent_imp_bounded:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3455
  fixes s :: "nat \<Rightarrow> 'a::metric_space"
44632
076a45f65e12 simplify/generalize some proofs
huffman
parents: 44628
diff changeset
  3456
  shows "(s ---> l) sequentially \<Longrightarrow> bounded (range s)"
50939
ae7cd20ed118 replace convergent_imp_cauchy by LIMSEQ_imp_Cauchy
hoelzl
parents: 50938
diff changeset
  3457
  by (intro cauchy_imp_bounded LIMSEQ_imp_Cauchy)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3458
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  3459
subsubsection {* Complete the chain of compactness variants *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3460
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3461
primrec helper_2::"(real \<Rightarrow> 'a::metric_space) \<Rightarrow> nat \<Rightarrow> 'a" where
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3462
  "helper_2 beyond 0 = beyond 0" |
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3463
  "helper_2 beyond (Suc n) = beyond (dist undefined (helper_2 beyond n) + 1 )"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3464
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3465
lemma bolzano_weierstrass_imp_bounded: fixes s::"'a::metric_space set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3466
  assumes "\<forall>t. infinite t \<and> t \<subseteq> s --> (\<exists>x \<in> s. x islimpt t)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3467
  shows "bounded s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3468
proof(rule ccontr)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3469
  assume "\<not> bounded s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3470
  then obtain beyond where "\<forall>a. beyond a \<in>s \<and> \<not> dist undefined (beyond a) \<le> a"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3471
    unfolding bounded_any_center [where a=undefined]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3472
    apply simp using choice[of "\<lambda>a x. x\<in>s \<and> \<not> dist undefined x \<le> a"] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3473
  hence beyond:"\<And>a. beyond a \<in>s" "\<And>a. dist undefined (beyond a) > a"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3474
    unfolding linorder_not_le by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3475
  def x \<equiv> "helper_2 beyond"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3476
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3477
  { fix m n ::nat assume "m<n"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3478
    hence "dist undefined (x m) + 1 < dist undefined (x n)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3479
    proof(induct n)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3480
      case 0 thus ?case by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3481
    next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3482
      case (Suc n)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3483
      have *:"dist undefined (x n) + 1 < dist undefined (x (Suc n))"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3484
        unfolding x_def and helper_2.simps
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3485
        using beyond(2)[of "dist undefined (helper_2 beyond n) + 1"] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3486
      thus ?case proof(cases "m < n")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3487
        case True thus ?thesis using Suc and * by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3488
      next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3489
        case False hence "m = n" using Suc(2) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3490
        thus ?thesis using * by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3491
      qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3492
    qed  } note * = this
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3493
  { fix m n ::nat assume "m\<noteq>n"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3494
    have "1 < dist (x m) (x n)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3495
    proof(cases "m<n")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3496
      case True
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3497
      hence "1 < dist undefined (x n) - dist undefined (x m)" using *[of m n] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3498
      thus ?thesis using dist_triangle [of undefined "x n" "x m"] by arith
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3499
    next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3500
      case False hence "n<m" using `m\<noteq>n` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3501
      hence "1 < dist undefined (x m) - dist undefined (x n)" using *[of n m] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3502
      thus ?thesis using dist_triangle2 [of undefined "x m" "x n"] by arith
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3503
    qed  } note ** = this
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3504
  { fix a b assume "x a = x b" "a \<noteq> b"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3505
    hence False using **[of a b] by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3506
  hence "inj x" unfolding inj_on_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3507
  moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3508
  { fix n::nat
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3509
    have "x n \<in> s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3510
    proof(cases "n = 0")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3511
      case True thus ?thesis unfolding x_def using beyond by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3512
    next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3513
      case False then obtain z where "n = Suc z" using not0_implies_Suc by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3514
      thus ?thesis unfolding x_def using beyond by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3515
    qed  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3516
  ultimately have "infinite (range x) \<and> range x \<subseteq> s" unfolding x_def using range_inj_infinite[of "helper_2 beyond"] using beyond(1) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3517
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3518
  then obtain l where "l\<in>s" and l:"l islimpt range x" using assms[THEN spec[where x="range x"]] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3519
  then obtain y where "x y \<noteq> l" and y:"dist (x y) l < 1/2" unfolding islimpt_approachable apply(erule_tac x="1/2" in allE) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3520
  then obtain z where "x z \<noteq> l" and z:"dist (x z) l < dist (x y) l" using l[unfolded islimpt_approachable, THEN spec[where x="dist (x y) l"]]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3521
    unfolding dist_nz by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3522
  show False using y and z and dist_triangle_half_l[of "x y" l 1 "x z"] and **[of y z] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3523
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3524
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  3525
text {* Hence express everything as an equivalence. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3526
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3527
lemma compact_eq_bolzano_weierstrass:
44074
e3a744a157d4 generalize compactness equivalence lemmas
huffman
parents: 44073
diff changeset
  3528
  fixes s :: "'a::metric_space set"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3529
  shows "compact s \<longleftrightarrow> (\<forall>t. infinite t \<and> t \<subseteq> s --> (\<exists>x \<in> s. x islimpt t))" (is "?lhs = ?rhs")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3530
proof
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  3531
  assume ?lhs thus ?rhs using heine_borel_imp_bolzano_weierstrass[of s] by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3532
next
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  3533
  assume ?rhs thus ?lhs
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  3534
    unfolding compact_eq_seq_compact_metric by (rule bolzano_weierstrass_imp_seq_compact)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3535
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3536
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3537
lemma nat_approx_posE:
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3538
  fixes e::real
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3539
  assumes "0 < e"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3540
  obtains n::nat where "1 / (Suc n) < e"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3541
proof atomize_elim
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3542
  have " 1 / real (Suc (nat (ceiling (1/e)))) < 1 / (ceiling (1/e))"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3543
    by (rule divide_strict_left_mono) (auto intro!: mult_pos_pos simp: `0 < e`)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3544
  also have "1 / (ceiling (1/e)) \<le> 1 / (1/e)"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3545
    by (rule divide_left_mono) (auto intro!: divide_pos_pos simp: `0 < e`)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3546
  also have "\<dots> = e" by simp
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3547
  finally show  "\<exists>n. 1 / real (Suc n) < e" ..
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3548
qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3549
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3550
lemma compact_eq_totally_bounded:
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  3551
  "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
  3552
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
  3553
  fix e::real
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3554
  def f \<equiv> "(\<lambda>x::'a. ball x e) ` UNIV"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3555
  assume "0 < e" "compact s"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3556
  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
  3557
    by (simp add: compact_eq_heine_borel)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3558
  moreover
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3559
  have d0: "\<And>x::'a. dist x x < e" using `0 < e` by simp
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3560
  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
  3561
  ultimately have "(\<exists>f'\<subseteq>f. finite f' \<and> s \<subseteq> \<Union>f')" ..
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3562
  then guess K .. note K = this
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3563
  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
  3564
  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
  3565
  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
  3566
    by (intro exI[where x="k ` K"]) (auto simp: f_def)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3567
next
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3568
  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
  3569
  show "compact s"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3570
  proof cases
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3571
    assume "s = {}" thus "compact s" by (simp add: compact_def)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3572
  next
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3573
    assume "s \<noteq> {}"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3574
    show ?thesis
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3575
      unfolding compact_def
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3576
    proof safe
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3577
      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
  3578
      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
  3579
      then obtain K where
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3580
        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
  3581
        unfolding choice_iff by blast
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3582
      {
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3583
        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
  3584
        assume "e > 0"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3585
        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
  3586
          by simp_all
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3587
        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
  3588
        proof (rule ccontr)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3589
          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
  3590
            using `s \<noteq> {}`
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3591
            by auto
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3592
          moreover
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3593
          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
  3594
          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
  3595
          ultimately
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3596
          show False using f'
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3597
          proof (induct arbitrary: s f f' rule: finite_ne_induct)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3598
            case (singleton x)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3599
            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
  3600
            thus ?case using singleton by (auto simp: ball_def)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3601
          next
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3602
            case (insert x A)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3603
            show ?case
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3604
            proof cases
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3605
              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
  3606
              have "infinite ((f o f') -` \<Union>((\<lambda>x. ball x e) ` (insert x A)))"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3607
                using insert by (intro infinite_super[OF _ inf_ms]) auto
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3608
              also have "((f o f') -` \<Union>((\<lambda>x. ball x e) ` (insert x A))) =
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3609
                {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
  3610
              finally have "infinite \<dots>" .
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3611
              moreover assume "finite {m. (f o f') m \<in> ball x e}"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3612
              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
  3613
              hence "A \<noteq> {}" by auto then obtain k where "k \<in> A" by auto
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3614
              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
  3615
              have r_mono: "\<And>n m. n < m \<Longrightarrow> r n < r m"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3616
                using enumerate_mono[OF _ inf] by (simp add: r_def)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3617
              hence "subseq r" by (simp add: subseq_def)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3618
              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
  3619
                using enumerate_in_set[OF inf] by (simp add: r_def)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3620
              show False
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3621
              proof (rule insert)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3622
                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
  3623
                fix k s assume "k \<in> A" "subseq s"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3624
                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
  3625
                  by (subst (2) o_assoc[symmetric]) (intro insert(6) subseq_o, simp_all)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3626
              next
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3627
                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
  3628
              qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3629
            next
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3630
              assume inf: "infinite {m. (f o f') m \<in> ball x e}"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3631
              def r \<equiv> "enumerate {m. (f o f') m \<in> ball x e}"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3632
              have r_mono: "\<And>n m. n < m \<Longrightarrow> r n < r m"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3633
                using enumerate_mono[OF _ inf] by (simp add: r_def)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3634
              hence "subseq r" by (simp add: subseq_def)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3635
              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
  3636
              moreover
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3637
              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
  3638
                using enumerate_in_set[OF inf] by (simp add: r_def)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3639
              hence "(f o f') (r i) \<in> ball x e" by simp
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3640
              ultimately show False by simp
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3641
            qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3642
          qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3643
        qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3644
      }
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3645
      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
  3646
      let ?e = "\<lambda>n. 1 / real (Suc n)"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3647
      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
  3648
      interpret subseqs ?P using ex by unfold_locales force
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3649
      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
  3650
        by (simp add: complete_def)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3651
      have "\<exists>l\<in>s. (f o diagseq) ----> l"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3652
      proof (intro limI metric_CauchyI)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3653
        fix e::real assume "0 < e" hence "0 < e / 2" by auto
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3654
        from nat_approx_posE[OF this] guess n . note n = this
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3655
        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
  3656
        proof (rule exI[where x="Suc n"], safe)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3657
          fix m mm assume "Suc n \<le> m" "Suc n \<le> mm"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3658
          let ?e = "1 / real (Suc n)"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3659
          from reducer_reduces[of n] obtain k where
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3660
            "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
  3661
            unfolding seqseq_reducer by auto
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3662
          moreover
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3663
          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
  3664
          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
  3665
          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
  3666
          finally
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3667
          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
  3668
            by (intro add_strict_mono) auto
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3669
          hence "dist ((f \<circ> diagseq) m) k + dist ((f \<circ> diagseq) mm) k < e"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3670
            by (simp add: dist_commute)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3671
          moreover have "dist ((f \<circ> diagseq) m) ((f \<circ> diagseq) mm) \<le>
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3672
            dist ((f \<circ> diagseq) m) k + dist ((f \<circ> diagseq) mm) k"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3673
            by (rule dist_triangle2)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3674
          ultimately show "dist ((f \<circ> diagseq) m) ((f \<circ> diagseq) mm) < e"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3675
            by simp
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3676
        qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3677
      next
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3678
        fix n show "(f o diagseq) n \<in> s" using f by simp
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3679
      qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3680
      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
  3681
    qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3682
  qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3683
qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3684
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3685
lemma compact_eq_bounded_closed:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3686
  fixes s :: "'a::heine_borel set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3687
  shows "compact s \<longleftrightarrow> bounded s \<and> closed s"  (is "?lhs = ?rhs")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3688
proof
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  3689
  assume ?lhs thus ?rhs
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  3690
    unfolding compact_eq_bolzano_weierstrass using bolzano_weierstrass_imp_bounded bolzano_weierstrass_imp_closed by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3691
next
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  3692
  assume ?rhs thus ?lhs
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  3693
    using bounded_closed_imp_seq_compact[of s] unfolding compact_eq_seq_compact_metric by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3694
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3695
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3696
lemma compact_imp_bounded:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3697
  fixes s :: "'a::metric_space set"
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  3698
  shows "compact s \<Longrightarrow> bounded s"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3699
proof -
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3700
  assume "compact s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3701
  hence "\<forall>t. infinite t \<and> t \<subseteq> s \<longrightarrow> (\<exists>x \<in> s. x islimpt t)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3702
    using heine_borel_imp_bolzano_weierstrass[of s] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3703
  thus "bounded s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3704
    by (rule bolzano_weierstrass_imp_bounded)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3705
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3706
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3707
lemma compact_cball[simp]:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3708
  fixes x :: "'a::heine_borel"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3709
  shows "compact(cball x e)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3710
  using compact_eq_bounded_closed bounded_cball closed_cball
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 compact_frontier_bounded[intro]:
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 "bounded s ==> compact(frontier s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3716
  unfolding frontier_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3717
  using compact_eq_bounded_closed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3718
  by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3719
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3720
lemma compact_frontier[intro]:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3721
  fixes s :: "'a::heine_borel set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3722
  shows "compact s ==> compact (frontier s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3723
  using compact_eq_bounded_closed compact_frontier_bounded
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3724
  by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3725
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3726
lemma frontier_subset_compact:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3727
  fixes s :: "'a::heine_borel set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3728
  shows "compact s ==> frontier s \<subseteq> s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3729
  using frontier_subset_closed compact_eq_bounded_closed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3730
  by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3731
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  3732
subsection {* Bounded closed nest property (proof does not use Heine-Borel) *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3733
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3734
lemma bounded_closed_nest:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3735
  assumes "\<forall>n. closed(s n)" "\<forall>n. (s n \<noteq> {})"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3736
  "(\<forall>m n. m \<le> n --> s n \<subseteq> s m)"  "bounded(s 0)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3737
  shows "\<exists>a::'a::heine_borel. \<forall>n::nat. a \<in> s(n)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3738
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3739
  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
  3740
  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
  3741
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3742
  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
  3743
    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
  3744
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3745
  { fix n::nat
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3746
    { fix e::real assume "e>0"
44907
93943da0a010 remove redundant lemma Lim_sequentially in favor of lemma LIMSEQ_def
huffman
parents: 44905
diff changeset
  3747
      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
  3748
      hence "dist ((x \<circ> r) (max N n)) l < e" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3749
      moreover
50937
d249ef928ae1 removed subseq_bigger (replaced by seq_suble)
hoelzl
parents: 50936
diff changeset
  3750
      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
  3751
      hence "(x \<circ> r) (max N n) \<in> s n"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3752
        using x apply(erule_tac x=n in allE)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3753
        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
  3754
        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
  3755
      ultimately have "\<exists>y\<in>s n. dist y l < e" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3756
    }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3757
    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
  3758
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3759
  thus ?thesis by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3760
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3761
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  3762
text {* Decreasing case does not even need compactness, just completeness. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3763
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3764
lemma decreasing_closed_nest:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3765
  assumes "\<forall>n. closed(s n)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3766
          "\<forall>n. (s n \<noteq> {})"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3767
          "\<forall>m n. m \<le> n --> s n \<subseteq> s m"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3768
          "\<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
  3769
  shows "\<exists>a::'a::complete_space. \<forall>n::nat. a \<in> s n"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3770
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3771
  have "\<forall>n. \<exists> x. x\<in>s n" using assms(2) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3772
  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
  3773
  then obtain t where t: "\<forall>n. t n \<in> s n" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3774
  { fix e::real assume "e>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3775
    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
  3776
    { fix m n ::nat assume "N \<le> m \<and> N \<le> n"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3777
      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
  3778
      hence "dist (t m) (t n) < e" using N by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3779
    }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3780
    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
  3781
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3782
  hence  "Cauchy t" unfolding cauchy_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3783
  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
  3784
  { fix n::nat
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3785
    { fix e::real assume "e>0"
44907
93943da0a010 remove redundant lemma Lim_sequentially in favor of lemma LIMSEQ_def
huffman
parents: 44905
diff changeset
  3786
      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
  3787
      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
  3788
      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
  3789
    }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3790
    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
  3791
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3792
  then show ?thesis by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3793
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3794
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  3795
text {* Strengthen it to the intersection actually being a singleton. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3796
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3797
lemma decreasing_closed_nest_sing:
44632
076a45f65e12 simplify/generalize some proofs
huffman
parents: 44628
diff changeset
  3798
  fixes s :: "nat \<Rightarrow> 'a::complete_space set"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3799
  assumes "\<forall>n. closed(s n)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3800
          "\<forall>n. s n \<noteq> {}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3801
          "\<forall>m n. m \<le> n --> s n \<subseteq> s m"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3802
          "\<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
  3803
  shows "\<exists>a. \<Inter>(range s) = {a}"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3804
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3805
  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
  3806
  { fix b assume b:"b \<in> \<Inter>(range s)"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3807
    { fix e::real assume "e>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3808
      hence "dist a b < e" using assms(4 )using b using a by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3809
    }
36778
739a9379e29b avoid using real-specific versions of generic lemmas
huffman
parents: 36669
diff changeset
  3810
    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
  3811
  }
34104
22758f95e624 re-state lemmas using 'range'
huffman
parents: 33758
diff changeset
  3812
  with a have "\<Inter>(range s) = {a}" unfolding image_def by auto
22758f95e624 re-state lemmas using 'range'
huffman
parents: 33758
diff changeset
  3813
  thus ?thesis ..
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3814
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3815
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3816
text{* Cauchy-type criteria for uniform convergence. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3817
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3818
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
  3819
 "(\<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
  3820
  (\<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
  3821
proof(rule)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3822
  assume ?lhs
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3823
  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
  3824
  { fix e::real assume "e>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3825
    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
  3826
    { 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
  3827
      hence "dist (s m x) (s n x) < e"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3828
        using N[THEN spec[where x=m], THEN spec[where x=x]]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3829
        using N[THEN spec[where x=n], THEN spec[where x=x]]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3830
        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
  3831
    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
  3832
  thus ?rhs by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3833
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3834
  assume ?rhs
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3835
  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
  3836
  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
  3837
    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
  3838
  { fix e::real assume "e>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3839
    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
  3840
      using `?rhs`[THEN spec[where x="e/2"]] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3841
    { fix x assume "P x"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3842
      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
  3843
        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
  3844
      fix n::nat assume "n\<ge>N"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3845
      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
  3846
        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
  3847
    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
  3848
  thus ?lhs by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3849
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3850
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3851
lemma uniformly_cauchy_imp_uniformly_convergent:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3852
  fixes s :: "nat \<Rightarrow> 'a \<Rightarrow> 'b::heine_borel"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3853
  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
  3854
          "\<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
  3855
  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
  3856
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3857
  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
  3858
    using assms(1) unfolding uniformly_convergent_eq_cauchy[THEN sym] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3859
  moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3860
  { fix x assume "P x"
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41969
diff changeset
  3861
    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
  3862
      using l and assms(2) unfolding LIMSEQ_def by blast  }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3863
  ultimately show ?thesis by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3864
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3865
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  3866
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  3867
subsection {* Continuity *}
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  3868
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  3869
text {* Define continuity over a net to take in restrictions of the set. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3870
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3871
definition
44081
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44076
diff changeset
  3872
  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
  3873
  where "continuous net f \<longleftrightarrow> (f ---> f(netlimit net)) net"
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_trivial_limit:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3876
 "trivial_limit net ==> continuous net f"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3877
  unfolding continuous_def tendsto_def trivial_limit_eq by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3878
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3879
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
  3880
  unfolding continuous_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3881
  unfolding tendsto_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3882
  using netlimit_within[of x s]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3883
  by (cases "trivial_limit (at x within s)") (auto simp add: trivial_limit_eventually)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3884
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3885
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
  3886
  using continuous_within [of x UNIV f] by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3887
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3888
lemma continuous_at_within:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3889
  assumes "continuous (at x) f"  shows "continuous (at x within s) f"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3890
  using assms unfolding continuous_at continuous_within
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3891
  by (rule Lim_at_within)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3892
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3893
text{* Derive the epsilon-delta forms, which we often use as "definitions" *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3894
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3895
lemma continuous_within_eps_delta:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3896
  "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
  3897
  unfolding continuous_within and Lim_within
44584
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  3898
  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
  3899
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3900
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
  3901
                           \<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
  3902
  using continuous_within_eps_delta [of x UNIV f] by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3903
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3904
text{* Versions in terms of open balls. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3905
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3906
lemma continuous_within_ball:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3907
 "continuous (at x within s) f \<longleftrightarrow> (\<forall>e>0. \<exists>d>0.
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3908
                            f ` (ball x d \<inter> s) \<subseteq> ball (f x) e)" (is "?lhs = ?rhs")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3909
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3910
  assume ?lhs
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3911
  { fix e::real assume "e>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3912
    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
  3913
      using `?lhs`[unfolded continuous_within Lim_within] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3914
    { fix y assume "y\<in>f ` (ball x d \<inter> s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3915
      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
  3916
        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
  3917
    }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3918
    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
  3919
  thus ?rhs by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3920
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3921
  assume ?rhs thus ?lhs unfolding continuous_within Lim_within ball_def subset_eq
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3922
    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
  3923
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3924
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3925
lemma continuous_at_ball:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3926
  "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
  3927
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3928
  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
  3929
    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
  3930
    unfolding dist_nz[THEN sym] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3931
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3932
  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
  3933
    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
  3934
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3935
36440
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3936
text{* Define setwise continuity in terms of limits within the set. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3937
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3938
definition
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  3939
  continuous_on ::
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  3940
    "'a set \<Rightarrow> ('a::topological_space \<Rightarrow> 'b::topological_space) \<Rightarrow> bool"
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  3941
where
36440
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3942
  "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
  3943
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3944
lemma continuous_on_topological:
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  3945
  "continuous_on s f \<longleftrightarrow>
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  3946
    (\<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
  3947
      (\<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
  3948
unfolding continuous_on_def tendsto_def
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3949
unfolding Limits.eventually_within eventually_at_topological
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3950
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
  3951
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  3952
lemma continuous_on_iff:
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  3953
  "continuous_on s f \<longleftrightarrow>
36440
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3954
    (\<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
  3955
unfolding continuous_on_def Lim_within
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3956
apply (intro ball_cong [OF refl] all_cong ex_cong)
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3957
apply (rename_tac y, case_tac "y = x", simp)
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3958
apply (simp add: dist_nz)
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  3959
done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3960
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3961
definition
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3962
  uniformly_continuous_on ::
36440
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3963
    "'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
  3964
where
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3965
  "uniformly_continuous_on s f \<longleftrightarrow>
36440
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3966
    (\<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
  3967
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3968
text{* Some simple consequential lemmas. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3969
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3970
lemma uniformly_continuous_imp_continuous:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3971
 " uniformly_continuous_on s f ==> continuous_on s f"
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  3972
  unfolding uniformly_continuous_on_def continuous_on_iff by blast
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3973
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3974
lemma continuous_at_imp_continuous_within:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3975
 "continuous (at x) f ==> continuous (at x within s) f"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3976
  unfolding continuous_within continuous_at using Lim_at_within by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3977
36440
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3978
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
  3979
unfolding tendsto_def by (simp add: trivial_limit_eq)
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3980
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  3981
lemma continuous_at_imp_continuous_on:
36440
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3982
  assumes "\<forall>x\<in>s. continuous (at x) f"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3983
  shows "continuous_on s f"
36440
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3984
unfolding continuous_on_def
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3985
proof
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3986
  fix x assume "x \<in> s"
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3987
  with assms have *: "(f ---> f (netlimit (at x))) (at x)"
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3988
    unfolding continuous_def by simp
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3989
  have "(f ---> f x) (at x)"
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3990
  proof (cases "trivial_limit (at x)")
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3991
    case True thus ?thesis
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3992
      by (rule Lim_trivial_limit)
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3993
  next
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3994
    case False
36667
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  3995
    hence 1: "netlimit (at x) = x"
45031
9583f2b56f85 add lemmas within_empty and tendsto_bot;
huffman
parents: 44909
diff changeset
  3996
      using netlimit_within [of x UNIV] by simp
36440
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3997
    with * show ?thesis by simp
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3998
  qed
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3999
  thus "(f ---> f x) (at x within s)"
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  4000
    by (rule Lim_at_within)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4001
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4002
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4003
lemma continuous_on_eq_continuous_within:
36440
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  4004
  "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
  4005
unfolding continuous_on_def continuous_def
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  4006
apply (rule ball_cong [OF refl])
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  4007
apply (case_tac "trivial_limit (at x within s)")
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  4008
apply (simp add: Lim_trivial_limit)
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  4009
apply (simp add: netlimit_within)
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  4010
done
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  4011
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  4012
lemmas continuous_on = continuous_on_def -- "legacy theorem name"
33175
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_eq_continuous_at:
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  4015
  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
  4016
  by (auto simp add: continuous_on continuous_at Lim_within_open)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4017
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4018
lemma continuous_within_subset:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4019
 "continuous (at x within s) f \<Longrightarrow> t \<subseteq> s
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4020
             ==> continuous (at x within t) f"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4021
  unfolding continuous_within by(metis Lim_within_subset)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4022
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4023
lemma continuous_on_subset:
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  4024
  shows "continuous_on s f \<Longrightarrow> t \<subseteq> s ==> continuous_on t f"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4025
  unfolding continuous_on by (metis subset_eq Lim_within_subset)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4026
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4027
lemma continuous_on_interior:
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  4028
  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
  4029
  by (erule interiorE, drule (1) continuous_on_subset,
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  4030
    simp add: continuous_on_eq_continuous_at)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4031
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4032
lemma continuous_on_eq:
36440
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  4033
  "(\<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
  4034
  unfolding continuous_on_def tendsto_def Limits.eventually_within
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  4035
  by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4036
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  4037
text {* Characterization of various kinds of continuity in terms of sequences. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4038
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4039
lemma continuous_within_sequentially:
44533
7abe4a59f75d generalize and simplify proof of continuous_within_sequentially
huffman
parents: 44531
diff changeset
  4040
  fixes f :: "'a::metric_space \<Rightarrow> 'b::topological_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4041
  shows "continuous (at a within s) f \<longleftrightarrow>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4042
                (\<forall>x. (\<forall>n::nat. x n \<in> s) \<and> (x ---> a) sequentially
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4043
                     --> ((f o x) ---> f a) sequentially)" (is "?lhs = ?rhs")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4044
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4045
  assume ?lhs
44533
7abe4a59f75d generalize and simplify proof of continuous_within_sequentially
huffman
parents: 44531
diff changeset
  4046
  { 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
  4047
    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
  4048
    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
  4049
      unfolding continuous_within tendsto_def eventually_within by auto
7abe4a59f75d generalize and simplify proof of continuous_within_sequentially
huffman
parents: 44531
diff changeset
  4050
    have "eventually (\<lambda>n. dist (x n) a < d) sequentially"
7abe4a59f75d generalize and simplify proof of continuous_within_sequentially
huffman
parents: 44531
diff changeset
  4051
      using x(2) `d>0` by simp
7abe4a59f75d generalize and simplify proof of continuous_within_sequentially
huffman
parents: 44531
diff changeset
  4052
    hence "eventually (\<lambda>n. (f \<circ> x) n \<in> T) sequentially"
46887
cb891d9a23c1 use eventually_elim method
noschinl
parents: 45776
diff changeset
  4053
    proof eventually_elim
cb891d9a23c1 use eventually_elim method
noschinl
parents: 45776
diff changeset
  4054
      case (elim n) thus ?case
44533
7abe4a59f75d generalize and simplify proof of continuous_within_sequentially
huffman
parents: 44531
diff changeset
  4055
        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
  4056
    qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4057
  }
44533
7abe4a59f75d generalize and simplify proof of continuous_within_sequentially
huffman
parents: 44531
diff changeset
  4058
  thus ?rhs unfolding tendsto_iff unfolding tendsto_def by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4059
next
44533
7abe4a59f75d generalize and simplify proof of continuous_within_sequentially
huffman
parents: 44531
diff changeset
  4060
  assume ?rhs thus ?lhs
7abe4a59f75d generalize and simplify proof of continuous_within_sequentially
huffman
parents: 44531
diff changeset
  4061
    unfolding continuous_within tendsto_def [where l="f a"]
7abe4a59f75d generalize and simplify proof of continuous_within_sequentially
huffman
parents: 44531
diff changeset
  4062
    by (simp add: sequentially_imp_eventually_within)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4063
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4064
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4065
lemma continuous_at_sequentially:
44533
7abe4a59f75d generalize and simplify proof of continuous_within_sequentially
huffman
parents: 44531
diff changeset
  4066
  fixes f :: "'a::metric_space \<Rightarrow> 'b::topological_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4067
  shows "continuous (at a) f \<longleftrightarrow> (\<forall>x. (x ---> a) sequentially
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4068
                  --> ((f o x) ---> f a) sequentially)"
45031
9583f2b56f85 add lemmas within_empty and tendsto_bot;
huffman
parents: 44909
diff changeset
  4069
  using continuous_within_sequentially[of a UNIV f] by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4070
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4071
lemma continuous_on_sequentially:
44533
7abe4a59f75d generalize and simplify proof of continuous_within_sequentially
huffman
parents: 44531
diff changeset
  4072
  fixes f :: "'a::metric_space \<Rightarrow> 'b::topological_space"
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  4073
  shows "continuous_on s f \<longleftrightarrow>
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  4074
    (\<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
  4075
                    --> ((f o x) ---> f(a)) sequentially)" (is "?lhs = ?rhs")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4076
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4077
  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
  4078
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4079
  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
  4080
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4081
44648
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4082
lemma uniformly_continuous_on_sequentially:
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4083
  "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
  4084
                    ((\<lambda>n. dist (x n) (y n)) ---> 0) sequentially
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4085
                    \<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
  4086
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4087
  assume ?lhs
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4088
  { 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
  4089
    { fix e::real assume "e>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4090
      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
  4091
        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
  4092
      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
  4093
      { fix n assume "n\<ge>N"
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4094
        hence "dist (f (x n)) (f (y n)) < e"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4095
          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
  4096
          unfolding dist_commute by simp  }
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4097
      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
  4098
    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
  4099
  thus ?rhs by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4100
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4101
  assume ?rhs
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4102
  { assume "\<not> ?lhs"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4103
    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
  4104
    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
  4105
      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
  4106
      by (auto simp add: dist_commute)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4107
    def x \<equiv> "\<lambda>n::nat. fst (fa (inverse (real n + 1)))"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4108
    def y \<equiv> "\<lambda>n::nat. snd (fa (inverse (real n + 1)))"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4109
    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
  4110
      unfolding x_def and y_def using fa by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4111
    { fix e::real assume "e>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4112
      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
  4113
      { fix n::nat assume "n\<ge>N"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4114
        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
  4115
        also have "\<dots> < e" using N by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4116
        finally have "inverse (real n + 1) < e" by auto
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4117
        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
  4118
      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
  4119
    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
  4120
    hence False using fxy and `e>0` by auto  }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4121
  thus ?lhs unfolding uniformly_continuous_on_def by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4122
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4123
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4124
text{* The usual transformation theorems. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4125
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4126
lemma continuous_transform_within:
36667
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  4127
  fixes f g :: "'a::metric_space \<Rightarrow> 'b::topological_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4128
  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
  4129
          "continuous (at x within s) f"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4130
  shows "continuous (at x within s) g"
36667
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  4131
unfolding continuous_within
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  4132
proof (rule Lim_transform_within)
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  4133
  show "0 < d" by fact
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  4134
  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
  4135
    using assms(3) by auto
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  4136
  have "f x = g x"
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  4137
    using assms(1,2,3) by auto
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  4138
  thus "(f ---> g x) (at x within s)"
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  4139
    using assms(4) unfolding continuous_within by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4140
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4141
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4142
lemma continuous_transform_at:
36667
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  4143
  fixes f g :: "'a::metric_space \<Rightarrow> 'b::topological_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4144
  assumes "0 < d" "\<forall>x'. dist x' x < d --> f x' = g x'"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4145
          "continuous (at x) f"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4146
  shows "continuous (at x) g"
45031
9583f2b56f85 add lemmas within_empty and tendsto_bot;
huffman
parents: 44909
diff changeset
  4147
  using continuous_transform_within [of d x UNIV f g] assms by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4148
44648
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4149
subsubsection {* Structural rules for pointwise continuity *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4150
44647
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4151
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
  4152
  unfolding continuous_within by (rule tendsto_ident_at_within)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4153
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4154
lemma continuous_at_id: "continuous (at a) (\<lambda>x. x)"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4155
  unfolding continuous_at by (rule tendsto_ident_at)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4156
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4157
lemma continuous_const: "continuous F (\<lambda>x. c)"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4158
  unfolding continuous_def by (rule tendsto_const)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4159
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4160
lemma continuous_dist:
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4161
  assumes "continuous F f" and "continuous F g"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4162
  shows "continuous F (\<lambda>x. dist (f x) (g x))"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4163
  using assms unfolding continuous_def by (rule tendsto_dist)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4164
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  4165
lemma continuous_infdist:
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  4166
  assumes "continuous F f"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  4167
  shows "continuous F (\<lambda>x. infdist (f x) A)"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  4168
  using assms unfolding continuous_def by (rule tendsto_infdist)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  4169
44647
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4170
lemma continuous_norm:
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4171
  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
  4172
  unfolding continuous_def by (rule tendsto_norm)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4173
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4174
lemma continuous_infnorm:
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4175
  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
  4176
  unfolding continuous_def by (rule tendsto_infnorm)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4177
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4178
lemma continuous_add:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4179
  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
  4180
  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
  4181
  unfolding continuous_def by (rule tendsto_add)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4182
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4183
lemma continuous_minus:
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4184
  fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4185
  shows "continuous F f \<Longrightarrow> continuous F (\<lambda>x. - f x)"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4186
  unfolding continuous_def by (rule tendsto_minus)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4187
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4188
lemma continuous_diff:
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4189
  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
  4190
  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
  4191
  unfolding continuous_def by (rule tendsto_diff)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4192
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4193
lemma continuous_scaleR:
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4194
  fixes g :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4195
  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
  4196
  unfolding continuous_def by (rule tendsto_scaleR)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4197
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4198
lemma continuous_mult:
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4199
  fixes f g :: "'a::t2_space \<Rightarrow> 'b::real_normed_algebra"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4200
  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
  4201
  unfolding continuous_def by (rule tendsto_mult)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4202
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4203
lemma continuous_inner:
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4204
  assumes "continuous F f" and "continuous F g"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4205
  shows "continuous F (\<lambda>x. inner (f x) (g x))"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4206
  using assms unfolding continuous_def by (rule tendsto_inner)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4207
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4208
lemma continuous_inverse:
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4209
  fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_div_algebra"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4210
  assumes "continuous F f" and "f (netlimit F) \<noteq> 0"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4211
  shows "continuous F (\<lambda>x. inverse (f x))"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4212
  using assms unfolding continuous_def by (rule tendsto_inverse)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4213
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4214
lemma continuous_at_within_inverse:
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4215
  fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_div_algebra"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4216
  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
  4217
  shows "continuous (at a within s) (\<lambda>x. inverse (f x))"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4218
  using assms unfolding continuous_within by (rule tendsto_inverse)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4219
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4220
lemma continuous_at_inverse:
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4221
  fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_div_algebra"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4222
  assumes "continuous (at a) f" and "f a \<noteq> 0"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4223
  shows "continuous (at a) (\<lambda>x. inverse (f x))"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4224
  using assms unfolding continuous_at by (rule tendsto_inverse)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4225
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4226
lemmas continuous_intros = continuous_at_id continuous_within_id
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4227
  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
  4228
  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
  4229
  continuous_inner continuous_at_inverse continuous_at_within_inverse
34964
4e8be3c04d37 Replaced vec1 and dest_vec1 by abbreviation.
hoelzl
parents: 34291
diff changeset
  4230
44648
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4231
subsubsection {* Structural rules for setwise continuity *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4232
44647
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4233
lemma continuous_on_id: "continuous_on s (\<lambda>x. x)"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4234
  unfolding continuous_on_def by (fast intro: tendsto_ident_at_within)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4235
44531
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4236
lemma continuous_on_const: "continuous_on s (\<lambda>x. c)"
44125
230a8665c919 mark some redundant theorems as legacy
huffman
parents: 44122
diff changeset
  4237
  unfolding continuous_on_def by (auto intro: tendsto_intros)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4238
44647
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4239
lemma continuous_on_norm:
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4240
  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
  4241
  unfolding continuous_on_def by (fast intro: tendsto_norm)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4242
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4243
lemma continuous_on_infnorm:
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4244
  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
  4245
  unfolding continuous_on by (fast intro: tendsto_infnorm)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4246
44531
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4247
lemma continuous_on_minus:
36440
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  4248
  fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4249
  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
  4250
  unfolding continuous_on_def by (auto intro: tendsto_intros)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4251
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4252
lemma continuous_on_add:
36440
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  4253
  fixes f g :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4254
  shows "continuous_on s f \<Longrightarrow> continuous_on s g
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4255
           \<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
  4256
  unfolding continuous_on_def by (auto intro: tendsto_intros)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4257
44531
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4258
lemma continuous_on_diff:
36440
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  4259
  fixes f g :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4260
  shows "continuous_on s f \<Longrightarrow> continuous_on s g
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4261
           \<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
  4262
  unfolding continuous_on_def by (auto intro: tendsto_intros)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4263
44531
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4264
lemma (in bounded_linear) continuous_on:
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4265
  "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
  4266
  unfolding continuous_on_def by (fast intro: tendsto)
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4267
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4268
lemma (in bounded_bilinear) continuous_on:
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4269
  "\<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
  4270
  unfolding continuous_on_def by (fast intro: tendsto)
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4271
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4272
lemma continuous_on_scaleR:
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4273
  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
  4274
  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
  4275
  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
  4276
  using bounded_bilinear_scaleR assms
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4277
  by (rule bounded_bilinear.continuous_on)
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4278
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4279
lemma continuous_on_mult:
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4280
  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
  4281
  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
  4282
  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
  4283
  using bounded_bilinear_mult assms
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4284
  by (rule bounded_bilinear.continuous_on)
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4285
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4286
lemma continuous_on_inner:
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4287
  fixes g :: "'a::topological_space \<Rightarrow> 'b::real_inner"
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4288
  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
  4289
  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
  4290
  using bounded_bilinear_inner assms
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4291
  by (rule bounded_bilinear.continuous_on)
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  4292
44647
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4293
lemma continuous_on_inverse:
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4294
  fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_div_algebra"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4295
  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
  4296
  shows "continuous_on s (\<lambda>x. inverse (f x))"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4297
  using assms unfolding continuous_on by (fast intro: tendsto_inverse)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4298
44648
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4299
subsubsection {* Structural rules for uniform continuity *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4300
44647
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4301
lemma uniformly_continuous_on_id:
44648
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4302
  shows "uniformly_continuous_on s (\<lambda>x. x)"
44647
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4303
  unfolding uniformly_continuous_on_def by auto
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4304
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4305
lemma uniformly_continuous_on_const:
44648
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4306
  shows "uniformly_continuous_on s (\<lambda>x. c)"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4307
  unfolding uniformly_continuous_on_def by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4308
44648
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4309
lemma uniformly_continuous_on_dist:
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4310
  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
  4311
  assumes "uniformly_continuous_on s f"
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4312
  assumes "uniformly_continuous_on s g"
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4313
  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
  4314
proof -
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4315
  { 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
  4316
      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
  4317
      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
  4318
      by arith
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4319
  } note le = this
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4320
  { fix x y
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4321
    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
  4322
    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
  4323
    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
  4324
      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
  4325
        simp add: le)
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4326
  }
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4327
  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
  4328
    unfolding dist_real_def by simp
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4329
qed
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4330
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4331
lemma uniformly_continuous_on_norm:
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4332
  assumes "uniformly_continuous_on s f"
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4333
  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
  4334
  unfolding norm_conv_dist using assms
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4335
  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
  4336
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4337
lemma (in bounded_linear) uniformly_continuous_on:
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4338
  assumes "uniformly_continuous_on s g"
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4339
  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
  4340
  using assms unfolding uniformly_continuous_on_sequentially
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4341
  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
  4342
  by (auto intro: tendsto_zero)
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4343
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4344
lemma uniformly_continuous_on_cmul:
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4345
  fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4346
  assumes "uniformly_continuous_on s f"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4347
  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
  4348
  using bounded_linear_scaleR_right assms
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4349
  by (rule bounded_linear.uniformly_continuous_on)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4350
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4351
lemma dist_minus:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4352
  fixes x y :: "'a::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4353
  shows "dist (- x) (- y) = dist x y"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4354
  unfolding dist_norm minus_diff_minus norm_minus_cancel ..
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4355
44648
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4356
lemma uniformly_continuous_on_minus:
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4357
  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
  4358
  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
  4359
  unfolding uniformly_continuous_on_def dist_minus .
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4360
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4361
lemma uniformly_continuous_on_add:
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4362
  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
  4363
  assumes "uniformly_continuous_on s f"
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4364
  assumes "uniformly_continuous_on s g"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4365
  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
  4366
  using assms unfolding uniformly_continuous_on_sequentially
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4367
  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
  4368
  by (auto intro: tendsto_add_zero)
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4369
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4370
lemma uniformly_continuous_on_diff:
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4371
  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
  4372
  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
  4373
  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
  4374
  unfolding ab_diff_minus using assms
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4375
  by (intro uniformly_continuous_on_add uniformly_continuous_on_minus)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4376
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4377
text{* Continuity of all kinds is preserved under composition. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4378
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4379
lemma continuous_within_topological:
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4380
  "continuous (at x within s) f \<longleftrightarrow>
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4381
    (\<forall>B. open B \<longrightarrow> f x \<in> B \<longrightarrow>
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4382
      (\<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
  4383
unfolding continuous_within
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4384
unfolding tendsto_def Limits.eventually_within eventually_at_topological
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4385
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
  4386
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4387
lemma continuous_within_compose:
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4388
  assumes "continuous (at x within s) f"
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4389
  assumes "continuous (at (f x) within f ` s) g"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4390
  shows "continuous (at x within s) (g o f)"
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4391
using assms unfolding continuous_within_topological by simp metis
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4392
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4393
lemma continuous_at_compose:
45031
9583f2b56f85 add lemmas within_empty and tendsto_bot;
huffman
parents: 44909
diff changeset
  4394
  assumes "continuous (at x) f" and "continuous (at (f x)) g"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4395
  shows "continuous (at x) (g o f)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4396
proof-
45031
9583f2b56f85 add lemmas within_empty and tendsto_bot;
huffman
parents: 44909
diff changeset
  4397
  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
  4398
    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
  4399
  thus ?thesis using assms(1)
9583f2b56f85 add lemmas within_empty and tendsto_bot;
huffman
parents: 44909
diff changeset
  4400
    using continuous_within_compose[of x UNIV f g] by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4401
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4402
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4403
lemma continuous_on_compose:
36440
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  4404
  "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
  4405
  unfolding continuous_on_topological by simp metis
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4406
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4407
lemma uniformly_continuous_on_compose:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4408
  assumes "uniformly_continuous_on s f"  "uniformly_continuous_on (f ` s) g"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4409
  shows "uniformly_continuous_on s (g o f)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4410
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4411
  { fix e::real assume "e>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4412
    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
  4413
    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
  4414
    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
  4415
  thus ?thesis using assms unfolding uniformly_continuous_on_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4416
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4417
44647
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4418
lemmas continuous_on_intros = continuous_on_id continuous_on_const
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4419
  continuous_on_compose continuous_on_norm continuous_on_infnorm
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4420
  continuous_on_add continuous_on_minus continuous_on_diff
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4421
  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
  4422
  continuous_on_inner
44648
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4423
  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
  4424
  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
  4425
  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
  4426
  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
  4427
  uniformly_continuous_on_cmul
44647
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4428
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4429
text{* Continuity in terms of open preimages. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4430
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4431
lemma continuous_at_open:
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4432
  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
  4433
unfolding continuous_within_topological [of x UNIV f, unfolded within_UNIV]
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4434
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
  4435
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4436
lemma continuous_on_open:
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4437
  shows "continuous_on s f \<longleftrightarrow>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4438
        (\<forall>t. openin (subtopology euclidean (f ` s)) t
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4439
            --> 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
  4440
proof (safe)
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4441
  fix t :: "'b set"
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4442
  assume 1: "continuous_on s f"
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4443
  assume 2: "openin (subtopology euclidean (f ` s)) t"
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4444
  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
  4445
    unfolding openin_open by auto
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4446
  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
  4447
  have "open U" unfolding U_def by (simp add: open_Union)
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4448
  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
  4449
  proof (intro ballI iffI)
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4450
    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
  4451
      unfolding U_def t by auto
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4452
  next
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4453
    fix x assume "x \<in> s" and "f x \<in> t"
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4454
    hence "x \<in> s" and "f x \<in> B"
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4455
      unfolding t by auto
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4456
    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
  4457
      unfolding t continuous_on_topological by metis
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4458
    then show "x \<in> U"
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4459
      unfolding U_def by auto
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4460
  qed
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4461
  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
  4462
  then show "openin (subtopology euclidean s) {x \<in> s. f x \<in> t}"
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4463
    unfolding openin_open by fast
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4464
next
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4465
  assume "?rhs" show "continuous_on s f"
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4466
  unfolding continuous_on_topological
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4467
  proof (clarify)
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4468
    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
  4469
    have "openin (subtopology euclidean (f ` s)) (f ` s \<inter> B)"
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4470
      unfolding openin_open using `open B` by auto
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4471
    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
  4472
      using `?rhs` by fast
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4473
    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
  4474
      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
  4475
  qed
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4476
qed
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4477
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4478
text {* Similarly in terms of closed sets. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4479
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4480
lemma continuous_on_closed:
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  4481
  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
  4482
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4483
  assume ?lhs
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4484
  { fix t
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4485
    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
  4486
    have **:"f ` s - (f ` s - (f ` s - t)) = f ` s - t" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4487
    assume as:"closedin (subtopology euclidean (f ` s)) t"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4488
    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
  4489
    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
  4490
      unfolding openin_closedin_eq topspace_euclidean_subtopology unfolding * by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4491
  thus ?rhs by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4492
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4493
  assume ?rhs
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4494
  { fix t
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4495
    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
  4496
    assume as:"openin (subtopology euclidean (f ` s)) t"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4497
    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
  4498
      unfolding openin_closedin_eq topspace_euclidean_subtopology *[THEN sym] closedin_subtopology by auto }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4499
  thus ?lhs unfolding continuous_on_open by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4500
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4501
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  4502
text {* Half-global and completely global cases. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4503
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4504
lemma continuous_open_in_preimage:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4505
  assumes "continuous_on s f"  "open t"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4506
  shows "openin (subtopology euclidean s) {x \<in> s. f x \<in> t}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4507
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4508
  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
  4509
  have "openin (subtopology euclidean (f ` s)) (t \<inter> f ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4510
    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
  4511
  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
  4512
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4513
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4514
lemma continuous_closed_in_preimage:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4515
  assumes "continuous_on s f"  "closed t"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4516
  shows "closedin (subtopology euclidean s) {x \<in> s. f x \<in> t}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4517
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4518
  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
  4519
  have "closedin (subtopology euclidean (f ` s)) (t \<inter> f ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4520
    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
  4521
  thus ?thesis
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4522
    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
  4523
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4524
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4525
lemma continuous_open_preimage:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4526
  assumes "continuous_on s f" "open s" "open t"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4527
  shows "open {x \<in> s. f x \<in> t}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4528
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4529
  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
  4530
    using continuous_open_in_preimage[OF assms(1,3)] unfolding openin_open by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4531
  thus ?thesis using open_Int[of s T, OF assms(2)] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4532
qed
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:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4535
  assumes "continuous_on s f" "closed s" "closed t"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4536
  shows "closed {x \<in> s. f x \<in> t}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4537
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4538
  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
  4539
    using continuous_closed_in_preimage[OF assms(1,3)] unfolding closedin_closed by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4540
  thus ?thesis using closed_Int[of s T, OF assms(2)] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4541
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4542
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4543
lemma continuous_open_preimage_univ:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4544
  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
  4545
  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
  4546
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4547
lemma continuous_closed_preimage_univ:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4548
  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
  4549
  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
  4550
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4551
lemma continuous_open_vimage:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4552
  shows "\<forall>x. continuous (at x) f \<Longrightarrow> open s \<Longrightarrow> open (f -` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4553
  unfolding vimage_def by (rule continuous_open_preimage_univ)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4554
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4555
lemma continuous_closed_vimage:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4556
  shows "\<forall>x. continuous (at x) f \<Longrightarrow> closed s \<Longrightarrow> closed (f -` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4557
  unfolding vimage_def by (rule continuous_closed_preimage_univ)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4558
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4559
lemma interior_image_subset:
35172
579dd5570f96 Added integration to Multivariate-Analysis (upto FTC)
himmelma
parents: 35028
diff changeset
  4560
  assumes "\<forall>x. continuous (at x) f" "inj f"
579dd5570f96 Added integration to Multivariate-Analysis (upto FTC)
himmelma
parents: 35028
diff changeset
  4561
  shows "interior (f ` s) \<subseteq> f ` (interior s)"
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  4562
proof
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  4563
  fix x assume "x \<in> interior (f ` s)"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  4564
  then obtain T where as: "open T" "x \<in> T" "T \<subseteq> f ` s" ..
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  4565
  hence "x \<in> f ` s" by auto
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  4566
  then obtain y where y: "y \<in> s" "x = f y" by auto
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  4567
  have "open (vimage f T)"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  4568
    using assms(1) `open T` by (rule continuous_open_vimage)
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  4569
  moreover have "y \<in> vimage f T"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  4570
    using `x = f y` `x \<in> T` by simp
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  4571
  moreover have "vimage f T \<subseteq> s"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  4572
    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
  4573
  ultimately have "y \<in> interior s" ..
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  4574
  with `x = f y` show "x \<in> f ` interior s" ..
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  4575
qed
35172
579dd5570f96 Added integration to Multivariate-Analysis (upto FTC)
himmelma
parents: 35028
diff changeset
  4576
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  4577
text {* Equality of continuous functions on closure and related results. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4578
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4579
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
  4580
  fixes f :: "_ \<Rightarrow> 'b::t1_space"
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  4581
  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
  4582
  using continuous_closed_in_preimage[of s f "{a}"] by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4583
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4584
lemma continuous_closed_preimage_constant:
36668
941ba2da372e simplify definition of t1_space; generalize lemma closed_sing and related lemmas
huffman
parents: 36667
diff changeset
  4585
  fixes f :: "_ \<Rightarrow> 'b::t1_space"
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  4586
  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
  4587
  using continuous_closed_preimage[of s f "{a}"] by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4588
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4589
lemma continuous_constant_on_closure:
36668
941ba2da372e simplify definition of t1_space; generalize lemma closed_sing and related lemmas
huffman
parents: 36667
diff changeset
  4590
  fixes f :: "_ \<Rightarrow> 'b::t1_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4591
  assumes "continuous_on (closure s) f"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4592
          "\<forall>x \<in> s. f x = a"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4593
  shows "\<forall>x \<in> (closure s). f x = a"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4594
    using continuous_closed_preimage_constant[of "closure s" f a]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4595
    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
  4596
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4597
lemma image_closure_subset:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4598
  assumes "continuous_on (closure s) f"  "closed t"  "(f ` s) \<subseteq> t"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4599
  shows "f ` (closure s) \<subseteq> t"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4600
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4601
  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
  4602
  moreover have "closed {x \<in> closure s. f x \<in> t}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4603
    using continuous_closed_preimage[OF assms(1)] and assms(2) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4604
  ultimately have "closure s = {x \<in> closure s . f x \<in> t}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4605
    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
  4606
  thus ?thesis by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4607
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4608
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4609
lemma continuous_on_closure_norm_le:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4610
  fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4611
  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
  4612
  shows "norm(f x) \<le> b"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4613
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4614
  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
  4615
  show ?thesis
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4616
    using image_closure_subset[OF assms(1) closed_cball[of 0 b] *] assms(3)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4617
    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
  4618
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4619
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  4620
text {* Making a continuous function avoid some value in a neighbourhood. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4621
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4622
lemma continuous_within_avoid:
50898
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4623
  fixes f :: "'a::metric_space \<Rightarrow> 'b::t1_space"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4624
  assumes "continuous (at x within s) f" and "f x \<noteq> a"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4625
  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
  4626
proof-
50898
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4627
  obtain U where "open U" and "f x \<in> U" and "a \<notin> U"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4628
    using t1_space [OF `f x \<noteq> a`] by fast
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4629
  have "(f ---> f x) (at x within s)"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4630
    using assms(1) by (simp add: continuous_within)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4631
  hence "eventually (\<lambda>y. f y \<in> U) (at x within s)"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4632
    using `open U` and `f x \<in> U`
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4633
    unfolding tendsto_def by fast
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4634
  hence "eventually (\<lambda>y. f y \<noteq> a) (at x within s)"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4635
    using `a \<notin> U` by (fast elim: eventually_mono [rotated])
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4636
  thus ?thesis
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4637
    unfolding Limits.eventually_within Limits.eventually_at
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4638
    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
  4639
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4640
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4641
lemma continuous_at_avoid:
50898
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4642
  fixes f :: "'a::metric_space \<Rightarrow> 'b::t1_space"
45031
9583f2b56f85 add lemmas within_empty and tendsto_bot;
huffman
parents: 44909
diff changeset
  4643
  assumes "continuous (at x) f" and "f x \<noteq> a"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4644
  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
  4645
  using assms continuous_within_avoid[of x UNIV f a] by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4646
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4647
lemma continuous_on_avoid:
50898
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4648
  fixes f :: "'a::metric_space \<Rightarrow> 'b::t1_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4649
  assumes "continuous_on s f"  "x \<in> s"  "f x \<noteq> a"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4650
  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
  4651
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
  4652
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4653
lemma continuous_on_open_avoid:
50898
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4654
  fixes f :: "'a::metric_space \<Rightarrow> 'b::t1_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4655
  assumes "continuous_on s f"  "open s"  "x \<in> s"  "f x \<noteq> a"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4656
  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
  4657
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
  4658
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  4659
text {* Proving a function is constant by proving open-ness of level set. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4660
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4661
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
  4662
  fixes f :: "_ \<Rightarrow> 'b::t1_space"
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  4663
  shows "connected s \<Longrightarrow> continuous_on s f \<Longrightarrow>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4664
        openin (subtopology euclidean s) {x \<in> s. f x = a}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4665
        ==> (\<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
  4666
unfolding connected_clopen using continuous_closed_in_preimage_constant by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4667
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4668
lemma continuous_levelset_open_in:
36668
941ba2da372e simplify definition of t1_space; generalize lemma closed_sing and related lemmas
huffman
parents: 36667
diff changeset
  4669
  fixes f :: "_ \<Rightarrow> 'b::t1_space"
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  4670
  shows "connected s \<Longrightarrow> continuous_on s f \<Longrightarrow>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4671
        openin (subtopology euclidean s) {x \<in> s. f x = a} \<Longrightarrow>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4672
        (\<exists>x \<in> s. f x = a)  ==> (\<forall>x \<in> s. f x = a)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4673
using continuous_levelset_open_in_cases[of s f ]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4674
by meson
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4675
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4676
lemma continuous_levelset_open:
36668
941ba2da372e simplify definition of t1_space; generalize lemma closed_sing and related lemmas
huffman
parents: 36667
diff changeset
  4677
  fixes f :: "_ \<Rightarrow> 'b::t1_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4678
  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
  4679
  shows "\<forall>x \<in> s. f x = a"
36362
06475a1547cb fix lots of looping simp calls and other warnings
huffman
parents: 36360
diff changeset
  4680
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
  4681
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  4682
text {* Some arithmetical combinations (more to prove). *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4683
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4684
lemma open_scaling[intro]:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4685
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4686
  assumes "c \<noteq> 0"  "open s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4687
  shows "open((\<lambda>x. c *\<^sub>R x) ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4688
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4689
  { fix x assume "x \<in> s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4690
    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
  4691
    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
  4692
    moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4693
    { fix y assume "dist y (c *\<^sub>R x) < e * \<bar>c\<bar>"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4694
      hence "norm ((1 / c) *\<^sub>R y - x) < e" unfolding dist_norm
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4695
        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
  4696
          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
  4697
      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
  4698
    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
  4699
  thus ?thesis unfolding open_dist by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4700
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4701
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4702
lemma minus_image_eq_vimage:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4703
  fixes A :: "'a::ab_group_add set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4704
  shows "(\<lambda>x. - x) ` A = (\<lambda>x. - x) -` A"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4705
  by (auto intro!: image_eqI [where f="\<lambda>x. - x"])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4706
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4707
lemma open_negations:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4708
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4709
  shows "open s ==> open ((\<lambda> x. -x) ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4710
  unfolding scaleR_minus1_left [symmetric]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4711
  by (rule open_scaling, auto)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4712
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4713
lemma open_translation:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4714
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4715
  assumes "open s"  shows "open((\<lambda>x. a + x) ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4716
proof-
44647
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4717
  { fix x have "continuous (at x) (\<lambda>x. x - a)"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4718
      by (intro continuous_diff continuous_at_id continuous_const) }
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4719
  moreover have "{x. x - a \<in> s} = op + a ` s" by force
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4720
  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
  4721
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4722
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4723
lemma open_affinity:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4724
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4725
  assumes "open s"  "c \<noteq> 0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4726
  shows "open ((\<lambda>x. a + c *\<^sub>R x) ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4727
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4728
  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
  4729
  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
  4730
  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
  4731
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4732
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4733
lemma interior_translation:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4734
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4735
  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
  4736
proof (rule set_eqI, rule)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4737
  fix x assume "x \<in> interior (op + a ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4738
  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
  4739
  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
  4740
  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
  4741
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4742
  fix x assume "x \<in> op + a ` interior s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4743
  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
  4744
  { fix z have *:"a + y - z = y + a - z" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4745
    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
  4746
    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
  4747
    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
  4748
  hence "ball x e \<subseteq> op + a ` s" unfolding subset_eq by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4749
  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
  4750
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4751
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4752
text {* Topological properties of linear functions. *}
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4753
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4754
lemma linear_lim_0:
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4755
  assumes "bounded_linear f" shows "(f ---> 0) (at (0))"
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4756
proof-
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4757
  interpret f: bounded_linear f by fact
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4758
  have "(f ---> f 0) (at 0)"
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4759
    using tendsto_ident_at by (rule f.tendsto)
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4760
  thus ?thesis unfolding f.zero .
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4761
qed
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4762
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4763
lemma linear_continuous_at:
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4764
  assumes "bounded_linear f"  shows "continuous (at a) f"
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4765
  unfolding continuous_at using assms
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4766
  apply (rule bounded_linear.tendsto)
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4767
  apply (rule tendsto_ident_at)
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4768
  done
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4769
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4770
lemma linear_continuous_within:
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4771
  shows "bounded_linear f ==> continuous (at x within s) f"
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4772
  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
  4773
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4774
lemma linear_continuous_on:
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4775
  shows "bounded_linear f ==> continuous_on s f"
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4776
  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
  4777
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  4778
text {* Also bilinear functions, in composition form. *}
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4779
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4780
lemma bilinear_continuous_at_compose:
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4781
  shows "continuous (at x) f \<Longrightarrow> continuous (at x) g \<Longrightarrow> bounded_bilinear h
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4782
        ==> continuous (at x) (\<lambda>x. h (f x) (g x))"
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4783
  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
  4784
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4785
lemma bilinear_continuous_within_compose:
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4786
  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
  4787
        ==> continuous (at x within s) (\<lambda>x. h (f x) (g x))"
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4788
  unfolding continuous_within using Lim_bilinear[of f "f x"] by auto
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4789
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4790
lemma bilinear_continuous_on_compose:
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4791
  shows "continuous_on s f \<Longrightarrow> continuous_on s g \<Longrightarrow> bounded_bilinear h
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4792
             ==> continuous_on s (\<lambda>x. h (f x) (g x))"
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4793
  unfolding continuous_on_def
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4794
  by (fast elim: bounded_bilinear.tendsto)
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4795
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  4796
text {* Preservation of compactness and connectedness under continuous function. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4797
50898
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4798
lemma compact_eq_openin_cover:
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4799
  "compact S \<longleftrightarrow>
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4800
    (\<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
proof safe
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4803
  fix C
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4804
  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
  4805
  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
  4806
    unfolding openin_open by force+
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4807
  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
  4808
    by (rule compactE)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4809
  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
  4810
    by auto
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4811
  thus "\<exists>D\<subseteq>C. finite D \<and> S \<subseteq> \<Union>D" ..
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4812
next
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4813
  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
  4814
        (\<exists>D\<subseteq>C. finite D \<and> S \<subseteq> \<Union>D)"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4815
  show "compact S"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4816
  proof (rule compactI)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4817
    fix C
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4818
    let ?C = "image (\<lambda>T. S \<inter> T) C"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4819
    assume "\<forall>t\<in>C. open t" and "S \<subseteq> \<Union>C"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4820
    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
  4821
      unfolding openin_open by auto
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4822
    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
  4823
      by metis
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4824
    let ?D = "inv_into C (\<lambda>T. S \<inter> T) ` D"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4825
    have "?D \<subseteq> C \<and> finite ?D \<and> S \<subseteq> \<Union>?D"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4826
    proof (intro conjI)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4827
      from `D \<subseteq> ?C` show "?D \<subseteq> C"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4828
        by (fast intro: inv_into_into)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4829
      from `finite D` show "finite ?D"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4830
        by (rule finite_imageI)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4831
      from `S \<subseteq> \<Union>D` show "S \<subseteq> \<Union>?D"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4832
        apply (rule subset_trans)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4833
        apply clarsimp
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4834
        apply (frule subsetD [OF `D \<subseteq> ?C`, THEN f_inv_into_f])
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4835
        apply (erule rev_bexI, fast)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4836
        done
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4837
    qed
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4838
    thus "\<exists>D\<subseteq>C. finite D \<and> S \<subseteq> \<Union>D" ..
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4839
  qed
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4840
qed
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4841
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4842
lemma compact_continuous_image:
50898
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4843
  assumes "continuous_on s f" and "compact s"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4844
  shows "compact (f ` s)"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4845
using assms (* FIXME: long unstructured proof *)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4846
unfolding continuous_on_open
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4847
unfolding compact_eq_openin_cover
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 (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
  4850
apply (drule mp)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4851
apply (rule conjI)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4852
apply simp
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4853
apply clarsimp
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4854
apply (drule subsetD)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4855
apply (erule imageI)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4856
apply fast
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4857
apply (erule thin_rl)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4858
apply clarify
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4859
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
  4860
apply (intro conjI)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4861
apply clarify
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4862
apply (rule inv_into_into)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4863
apply (erule (1) subsetD)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4864
apply (erule finite_imageI)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4865
apply (clarsimp, rename_tac x)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4866
apply (drule (1) subsetD, clarify)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4867
apply (drule (1) subsetD, clarify)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4868
apply (rule rev_bexI)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4869
apply assumption
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4870
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
  4871
apply (drule f_inv_into_f)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4872
apply fast
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4873
apply (erule imageI)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4874
done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4875
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4876
lemma connected_continuous_image:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4877
  assumes "continuous_on s f"  "connected s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4878
  shows "connected(f ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4879
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4880
  { 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
  4881
    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
  4882
      using assms(1)[unfolded continuous_on_open, THEN spec[where x=T]]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4883
      using assms(1)[unfolded continuous_on_closed, THEN spec[where x=T]]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4884
      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
  4885
    hence False using as(1,2)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4886
      using as(4)[unfolded closedin_def topspace_euclidean_subtopology] by auto }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4887
  thus ?thesis unfolding connected_clopen by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4888
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4889
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  4890
text {* Continuity implies uniform continuity on a compact domain. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4891
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4892
lemma compact_uniformly_continuous:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4893
  assumes "continuous_on s f"  "compact s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4894
  shows "uniformly_continuous_on s f"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4895
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4896
    { fix x assume x:"x\<in>s"
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  4897
      hence "\<forall>xa. \<exists>y. 0 < xa \<longrightarrow> (y > 0 \<and> (\<forall>x'\<in>s. dist x' x < y \<longrightarrow> dist (f x') (f x) < xa))" using assms(1)[unfolded continuous_on_iff, THEN bspec[where x=x]] by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4898
      hence "\<exists>fa. \<forall>xa>0. \<forall>x'\<in>s. fa xa > 0 \<and> (dist x' x < fa xa \<longrightarrow> dist (f x') (f x) < xa)" using choice[of "\<lambda>e d. e>0 \<longrightarrow> d>0 \<and>(\<forall>x'\<in>s. (dist x' x < d \<longrightarrow> dist (f x') (f x) < e))"] by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4899
    then have "\<forall>x\<in>s. \<exists>y. \<forall>xa. 0 < xa \<longrightarrow> (\<forall>x'\<in>s. y xa > 0 \<and> (dist x' x < y xa \<longrightarrow> dist (f x') (f x) < xa))" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4900
    then obtain d where d:"\<forall>e>0. \<forall>x\<in>s. \<forall>x'\<in>s. d x e > 0 \<and> (dist x' x < d x e \<longrightarrow> dist (f x') (f x) < e)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4901
      using bchoice[of s "\<lambda>x fa. \<forall>xa>0. \<forall>x'\<in>s. fa xa > 0 \<and> (dist x' x < fa xa \<longrightarrow> dist (f x') (f x) < xa)"] by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4902
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4903
  { fix e::real assume "e>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4904
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4905
    { fix x assume "x\<in>s" hence "x \<in> ball x (d x (e / 2))" unfolding centre_in_ball using d[THEN spec[where x="e/2"]] using `e>0` by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4906
    hence "s \<subseteq> \<Union>{ball x (d x (e / 2)) |x. x \<in> s}" unfolding subset_eq by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4907
    moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4908
    { fix b assume "b\<in>{ball x (d x (e / 2)) |x. x \<in> s}" hence "open b" by auto  }
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  4909
    ultimately obtain ea where "ea>0" and ea:"\<forall>x\<in>s. \<exists>b\<in>{ball x (d x (e / 2)) |x. x \<in> s}. ball x ea \<subseteq> b"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  4910
      using heine_borel_lemma[OF assms(2)[unfolded compact_eq_seq_compact_metric], of "{ball x (d x (e / 2)) | x. x\<in>s }"] by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4911
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4912
    { fix x y assume "x\<in>s" "y\<in>s" and as:"dist y x < ea"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4913
      obtain z where "z\<in>s" and z:"ball x ea \<subseteq> ball z (d z (e / 2))" using ea[THEN bspec[where x=x]] and `x\<in>s` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4914
      hence "x\<in>ball z (d z (e / 2))" using `ea>0` unfolding subset_eq by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4915
      hence "dist (f z) (f x) < e / 2" using d[THEN spec[where x="e/2"]] and `e>0` and `x\<in>s` and `z\<in>s`
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4916
        by (auto  simp add: dist_commute)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4917
      moreover have "y\<in>ball z (d z (e / 2))" using as and `ea>0` and z[unfolded subset_eq]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4918
        by (auto simp add: dist_commute)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4919
      hence "dist (f z) (f y) < e / 2" using d[THEN spec[where x="e/2"]] and `e>0` and `y\<in>s` and `z\<in>s`
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4920
        by (auto  simp add: dist_commute)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4921
      ultimately have "dist (f y) (f x) < e" using dist_triangle_half_r[of "f z" "f x" e "f y"]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4922
        by (auto simp add: dist_commute)  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4923
    then have "\<exists>d>0. \<forall>x\<in>s. \<forall>x'\<in>s. dist x' x < d \<longrightarrow> dist (f x') (f x) < e" using `ea>0` by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4924
  thus ?thesis unfolding uniformly_continuous_on_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4925
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4926
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4927
text{* Continuity of inverse function on compact domain. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4928
44647
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4929
lemma continuous_on_inv:
50898
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4930
  fixes f :: "'a::topological_space \<Rightarrow> 'b::t2_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4931
  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
  4932
  shows "continuous_on (f ` s) g"
50898
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4933
unfolding continuous_on_topological
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4934
proof (clarsimp simp add: assms(3))
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4935
  fix x :: 'a and B :: "'a set"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4936
  assume "x \<in> s" and "open B" and "x \<in> B"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4937
  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
  4938
    using assms(3) by (auto, metis)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4939
  have "continuous_on (s - B) f"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4940
    using `continuous_on s f` Diff_subset
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4941
    by (rule continuous_on_subset)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4942
  moreover have "compact (s - B)"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4943
    using `open B` and `compact s`
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4944
    unfolding Diff_eq by (intro compact_inter_closed closed_Compl)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4945
  ultimately have "compact (f ` (s - B))"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4946
    by (rule compact_continuous_image)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4947
  hence "closed (f ` (s - B))"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4948
    by (rule compact_imp_closed)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4949
  hence "open (- f ` (s - B))"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4950
    by (rule open_Compl)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4951
  moreover have "f x \<in> - f ` (s - B)"
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4952
    using `x \<in> s` and `x \<in> B` by (simp add: 1)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4953
  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
  4954
    by (simp add: 1)
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  4955
  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
  4956
    by fast
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4957
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4958
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4959
text {* A uniformly convergent limit of continuous functions is continuous. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4960
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4961
lemma continuous_uniform_limit:
44212
4d10e7f342b1 generalize lemma continuous_uniform_limit to class metric_space
huffman
parents: 44211
diff changeset
  4962
  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
  4963
  assumes "\<not> trivial_limit F"
4d10e7f342b1 generalize lemma continuous_uniform_limit to class metric_space
huffman
parents: 44211
diff changeset
  4964
  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
  4965
  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
  4966
  shows "continuous_on s g"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4967
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4968
  { 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
  4969
    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
  4970
      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
  4971
    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
  4972
    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
  4973
      using assms(1) by blast
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4974
    have "e / 3 > 0" using `e>0` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4975
    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
  4976
      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
  4977
    { 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
  4978
      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
  4979
        by (rule d [rule_format])
4d10e7f342b1 generalize lemma continuous_uniform_limit to class metric_space
huffman
parents: 44211
diff changeset
  4980
      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
  4981
        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
  4982
        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
  4983
        by auto
4d10e7f342b1 generalize lemma continuous_uniform_limit to class metric_space
huffman
parents: 44211
diff changeset
  4984
      hence "dist (g y) (g x) < e"
4d10e7f342b1 generalize lemma continuous_uniform_limit to class metric_space
huffman
parents: 44211
diff changeset
  4985
        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
  4986
        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
  4987
        by auto }
4d10e7f342b1 generalize lemma continuous_uniform_limit to class metric_space
huffman
parents: 44211
diff changeset
  4988
    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
  4989
      using `d>0` by auto }
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  4990
  thus ?thesis unfolding continuous_on_iff by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4991
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4992
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  4993
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  4994
subsection {* Topological stuff lifted from and dropped to R *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4995
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4996
lemma open_real:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4997
  fixes s :: "real set" shows
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4998
 "open s \<longleftrightarrow>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4999
        (\<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
  5000
  unfolding open_dist 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 islimpt_approachable_real:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5003
  fixes s :: "real set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5004
  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
  5005
  unfolding islimpt_approachable dist_norm by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5006
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5007
lemma closed_real:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5008
  fixes s :: "real set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5009
  shows "closed s \<longleftrightarrow>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5010
        (\<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
  5011
            --> x \<in> s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5012
  unfolding closed_limpt islimpt_approachable dist_norm by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5013
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5014
lemma continuous_at_real_range:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5015
  fixes f :: "'a::real_normed_vector \<Rightarrow> real"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5016
  shows "continuous (at x) f \<longleftrightarrow> (\<forall>e>0. \<exists>d>0.
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5017
        \<forall>x'. norm(x' - x) < d --> abs(f x' - f x) < e)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5018
  unfolding continuous_at unfolding Lim_at
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5019
  unfolding dist_nz[THEN sym] unfolding dist_norm apply auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5020
  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
  5021
  apply(erule_tac x=e in allE) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5022
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5023
lemma continuous_on_real_range:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5024
  fixes f :: "'a::real_normed_vector \<Rightarrow> real"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5025
  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
  5026
  unfolding continuous_on_iff dist_norm by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5027
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  5028
text {* Hence some handy theorems on distance, diameter etc. of/from a set. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5029
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5030
lemma compact_attains_sup:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5031
  fixes s :: "real set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5032
  assumes "compact s"  "s \<noteq> {}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5033
  shows "\<exists>x \<in> s. \<forall>y \<in> s. y \<le> x"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5034
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5035
  from assms(1) have a:"bounded s" "closed s" unfolding compact_eq_bounded_closed by auto
33270
paulson
parents: 33175
diff changeset
  5036
  { 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
  5037
    have "isLub UNIV s (Sup s)" using Sup[OF assms(2)] unfolding setle_def using as(1) by auto
paulson
parents: 33175
diff changeset
  5038
    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
  5039
    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
  5040
  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
  5041
    apply(rule_tac x="Sup s" in bexI) by auto
paulson
parents: 33175
diff changeset
  5042
qed
paulson
parents: 33175
diff changeset
  5043
paulson
parents: 33175
diff changeset
  5044
lemma Inf:
paulson
parents: 33175
diff changeset
  5045
  fixes S :: "real set"
paulson
parents: 33175
diff changeset
  5046
  shows "S \<noteq> {} ==> (\<exists>b. b <=* S) ==> isGlb UNIV S (Inf S)"
paulson
parents: 33175
diff changeset
  5047
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
  5048
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5049
lemma compact_attains_inf:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5050
  fixes s :: "real set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5051
  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
  5052
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5053
  from assms(1) have a:"bounded s" "closed s" unfolding compact_eq_bounded_closed by auto
33270
paulson
parents: 33175
diff changeset
  5054
  { fix e::real assume as: "\<forall>x\<in>s. x \<ge> Inf s"  "Inf s \<notin> s"  "0 < e"
paulson
parents: 33175
diff changeset
  5055
      "\<forall>x'\<in>s. x' = Inf s \<or> \<not> abs (x' - Inf s) < e"
paulson
parents: 33175
diff changeset
  5056
    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
  5057
    moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5058
    { fix x assume "x \<in> s"
33270
paulson
parents: 33175
diff changeset
  5059
      hence *:"abs (x - Inf s) = x - Inf s" using as(1)[THEN bspec[where x=x]] by auto
paulson
parents: 33175
diff changeset
  5060
      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
  5061
    hence "isLb UNIV s (Inf s + e)" unfolding isLb_def and setge_def by auto
paulson
parents: 33175
diff changeset
  5062
    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
  5063
  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
  5064
    apply(rule_tac x="Inf s" in bexI) by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5065
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5066
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5067
lemma continuous_attains_sup:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5068
  fixes f :: "'a::metric_space \<Rightarrow> real"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5069
  shows "compact s \<Longrightarrow> s \<noteq> {} \<Longrightarrow> continuous_on s f
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5070
        ==> (\<exists>x \<in> s. \<forall>y \<in> s.  f y \<le> f x)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5071
  using compact_attains_sup[of "f ` s"]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5072
  using compact_continuous_image[of s f] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5073
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5074
lemma continuous_attains_inf:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5075
  fixes f :: "'a::metric_space \<Rightarrow> real"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5076
  shows "compact s \<Longrightarrow> s \<noteq> {} \<Longrightarrow> continuous_on s f
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5077
        \<Longrightarrow> (\<exists>x \<in> s. \<forall>y \<in> s. f x \<le> f y)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5078
  using compact_attains_inf[of "f ` s"]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5079
  using compact_continuous_image[of s f] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5080
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5081
lemma distance_attains_sup:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5082
  assumes "compact s" "s \<noteq> {}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5083
  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
  5084
proof (rule continuous_attains_sup [OF assms])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5085
  { fix x assume "x\<in>s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5086
    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
  5087
      by (intro tendsto_dist tendsto_const Lim_at_within tendsto_ident_at)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5088
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5089
  thus "continuous_on s (dist a)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5090
    unfolding continuous_on ..
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5091
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5092
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  5093
text {* For \emph{minimal} distance, we only need closure, not compactness. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5094
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5095
lemma distance_attains_inf:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5096
  fixes a :: "'a::heine_borel"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5097
  assumes "closed s"  "s \<noteq> {}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5098
  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
  5099
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5100
  from assms(2) obtain b where "b\<in>s" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5101
  let ?B = "cball a (dist b a) \<inter> s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5102
  have "b \<in> ?B" using `b\<in>s` by (simp add: dist_commute)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5103
  hence "?B \<noteq> {}" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5104
  moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5105
  { fix x assume "x\<in>?B"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5106
    fix e::real assume "e>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5107
    { fix x' assume "x'\<in>?B" and as:"dist x' x < e"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5108
      from as have "\<bar>dist a x' - dist a x\<bar> < e"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5109
        unfolding abs_less_iff minus_diff_eq
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5110
        using dist_triangle2 [of a x' x]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5111
        using dist_triangle [of a x x']
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5112
        by arith
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5113
    }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5114
    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
  5115
      using `e>0` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5116
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5117
  hence "continuous_on (cball a (dist b a) \<inter> s) (dist a)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5118
    unfolding continuous_on Lim_within dist_norm real_norm_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5119
    by fast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5120
  moreover have "compact ?B"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5121
    using compact_cball[of a "dist b a"]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5122
    unfolding compact_eq_bounded_closed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5123
    using bounded_Int and closed_Int and assms(1) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5124
  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
  5125
    using continuous_attains_inf[of ?B "dist a"] by fastforce
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 44668
diff changeset
  5126
  thus ?thesis by fastforce
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5127
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5128
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  5129
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  5130
subsection {* Pasted sets *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5131
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5132
lemma bounded_Times:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5133
  assumes "bounded s" "bounded t" shows "bounded (s \<times> t)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5134
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5135
  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
  5136
    using assms [unfolded bounded_def] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5137
  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
  5138
    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
  5139
  thus ?thesis unfolding bounded_any_center [where a="(x, y)"] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5140
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5141
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5142
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
  5143
by (induct x) simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5144
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  5145
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
  5146
unfolding seq_compact_def
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5147
apply clarify
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5148
apply (drule_tac x="fst \<circ> f" in spec)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5149
apply (drule mp, simp add: mem_Times_iff)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5150
apply (clarify, rename_tac l1 r1)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5151
apply (drule_tac x="snd \<circ> f \<circ> r1" in spec)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5152
apply (drule mp, simp add: mem_Times_iff)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5153
apply (clarify, rename_tac l2 r2)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5154
apply (rule_tac x="(l1, l2)" in rev_bexI, simp)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5155
apply (rule_tac x="r1 \<circ> r2" in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5156
apply (rule conjI, simp add: subseq_def)
48125
602dc0215954 tuned proofs -- prefer direct "rotated" instead of old-style COMP;
wenzelm
parents: 48048
diff changeset
  5157
apply (drule_tac r=r2 in lim_subseq [rotated], assumption)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5158
apply (drule (1) tendsto_Pair) back
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5159
apply (simp add: o_def)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5160
done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5161
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  5162
text {* Generalize to @{class topological_space} *}
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  5163
lemma compact_Times: 
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  5164
  fixes s :: "'a::metric_space set" and t :: "'b::metric_space set"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  5165
  shows "compact s \<Longrightarrow> compact t \<Longrightarrow> compact (s \<times> t)"
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  5166
  unfolding compact_eq_seq_compact_metric by (rule seq_compact_Times)
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50883
diff changeset
  5167
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  5168
text{* Hence some useful properties follow quite easily. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5169
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5170
lemma compact_scaling:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5171
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5172
  assumes "compact s"  shows "compact ((\<lambda>x. c *\<^sub>R x) ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5173
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5174
  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
  5175
  have *:"bounded_linear ?f" by (rule bounded_linear_scaleR_right)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5176
  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
  5177
    using linear_continuous_at[OF *] assms by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5178
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5179
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5180
lemma compact_negations:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5181
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5182
  assumes "compact s"  shows "compact ((\<lambda>x. -x) ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5183
  using compact_scaling [OF assms, of "- 1"] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5184
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5185
lemma compact_sums:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5186
  fixes s t :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5187
  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
  5188
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5189
  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
  5190
    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
  5191
  have "continuous_on (s \<times> t) (\<lambda>z. fst z + snd z)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5192
    unfolding continuous_on by (rule ballI) (intro tendsto_intros)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5193
  thus ?thesis unfolding * using compact_continuous_image compact_Times [OF assms] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5194
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5195
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5196
lemma compact_differences:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5197
  fixes s t :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5198
  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
  5199
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5200
  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
  5201
    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
  5202
  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
  5203
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5204
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5205
lemma compact_translation:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5206
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5207
  assumes "compact s"  shows "compact ((\<lambda>x. a + x) ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5208
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5209
  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
  5210
  thus ?thesis using compact_sums[OF assms compact_sing[of a]] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5211
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5212
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5213
lemma compact_affinity:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5214
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5215
  assumes "compact s"  shows "compact ((\<lambda>x. a + c *\<^sub>R x) ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5216
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5217
  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
  5218
  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
  5219
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5220
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  5221
text {* Hence we get the following. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5222
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5223
lemma compact_sup_maxdistance:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5224
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5225
  assumes "compact s"  "s \<noteq> {}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5226
  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
  5227
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5228
  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
  5229
  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
  5230
    using compact_differences[OF assms(1) assms(1)]
36362
06475a1547cb fix lots of looping simp calls and other warnings
huffman
parents: 36360
diff changeset
  5231
    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
  5232
  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
  5233
  thus ?thesis using x(2)[unfolded `x = a - b`] by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5234
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5235
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  5236
text {* We can state this in terms of diameter of a set. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5237
33270
paulson
parents: 33175
diff changeset
  5238
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
  5239
  (* TODO: generalize to class metric_space *)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5240
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5241
lemma diameter_bounded:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5242
  assumes "bounded s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5243
  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
  5244
        "\<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
  5245
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5246
  let ?D = "{norm (x - y) |x y. x \<in> s \<and> y \<in> s}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5247
  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
  5248
  { 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
  5249
    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
  5250
  note * = this
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5251
  { 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
  5252
    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
  5253
      by simp (blast del: Sup_upper intro!: * Sup_upper) }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5254
  moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5255
  { fix d::real assume "d>0" "d < diameter s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5256
    hence "s\<noteq>{}" unfolding diameter_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5257
    have "\<exists>d' \<in> ?D. d' > d"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5258
    proof(rule ccontr)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5259
      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
  5260
      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
  5261
      thus False using `d < diameter s` `s\<noteq>{}` 
51eb2ffa2189 Tidied up some very ugly proofs
paulson
parents: 33270
diff changeset
  5262
        apply (auto simp add: diameter_def) 
51eb2ffa2189 Tidied up some very ugly proofs
paulson
parents: 33270
diff changeset
  5263
        apply (drule Sup_real_iff [THEN [2] rev_iffD2])
51eb2ffa2189 Tidied up some very ugly proofs
paulson
parents: 33270
diff changeset
  5264
        apply (auto, force) 
51eb2ffa2189 Tidied up some very ugly proofs
paulson
parents: 33270
diff changeset
  5265
        done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5266
    qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5267
    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
  5268
  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
  5269
        "\<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
  5270
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5271
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5272
lemma diameter_bounded_bound:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5273
 "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
  5274
  using diameter_bounded by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5275
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5276
lemma diameter_compact_attained:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5277
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5278
  assumes "compact s"  "s \<noteq> {}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5279
  shows "\<exists>x\<in>s. \<exists>y\<in>s. (norm(x - y) = diameter s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5280
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5281
  have b:"bounded s" using assms(1) by (rule compact_imp_bounded)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5282
  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
  5283
  hence "diameter s \<le> norm (x - y)"
06475a1547cb fix lots of looping simp calls and other warnings
huffman
parents: 36360
diff changeset
  5284
    unfolding diameter_def by clarsimp (rule Sup_least, fast+)
33324
51eb2ffa2189 Tidied up some very ugly proofs
paulson
parents: 33270
diff changeset
  5285
  thus ?thesis
36362
06475a1547cb fix lots of looping simp calls and other warnings
huffman
parents: 36360
diff changeset
  5286
    by (metis b diameter_bounded_bound order_antisym xys)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5287
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5288
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  5289
text {* Related results with closure as the conclusion. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5290
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5291
lemma closed_scaling:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5292
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5293
  assumes "closed s" shows "closed ((\<lambda>x. c *\<^sub>R x) ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5294
proof(cases "s={}")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5295
  case True thus ?thesis by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5296
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5297
  case False
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5298
  show ?thesis
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5299
  proof(cases "c=0")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5300
    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
  5301
    case True thus ?thesis apply auto unfolding * by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5302
  next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5303
    case False
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5304
    { 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
  5305
      { fix n::nat have "scaleR (1 / c) (x n) \<in> s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5306
          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
  5307
          using `c\<noteq>0` by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5308
      }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5309
      moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5310
      { fix e::real assume "e>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5311
        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
  5312
        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
  5313
          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
  5314
        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
  5315
          unfolding dist_norm unfolding scaleR_right_diff_distrib[THEN sym]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5316
          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
  5317
      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
  5318
      ultimately have "l \<in> scaleR c ` s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5319
        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
  5320
        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
  5321
    thus ?thesis unfolding closed_sequential_limits by fast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5322
  qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5323
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5324
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5325
lemma closed_negations:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5326
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5327
  assumes "closed s"  shows "closed ((\<lambda>x. -x) ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5328
  using closed_scaling[OF assms, of "- 1"] by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5329
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5330
lemma compact_closed_sums:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5331
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5332
  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
  5333
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5334
  let ?S = "{x + y |x y. x \<in> s \<and> y \<in> t}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5335
  { fix x l assume as:"\<forall>n. x n \<in> ?S"  "(x ---> l) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5336
    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
  5337
      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
  5338
    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
  5339
      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
  5340
    have "((\<lambda>n. snd (f (r n))) ---> l - l') sequentially"
44125
230a8665c919 mark some redundant theorems as legacy
huffman
parents: 44122
diff changeset
  5341
      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
  5342
    hence "l - l' \<in> t"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5343
      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
  5344
      using f(3) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5345
    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
  5346
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5347
  thus ?thesis unfolding closed_sequential_limits by fast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5348
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5349
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5350
lemma closed_compact_sums:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5351
  fixes s t :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5352
  assumes "closed s"  "compact t"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5353
  shows "closed {x + y | x y. x \<in> s \<and> y \<in> t}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5354
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5355
  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
  5356
    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
  5357
  thus ?thesis using compact_closed_sums[OF assms(2,1)] by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5358
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5359
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5360
lemma compact_closed_differences:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5361
  fixes s t :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5362
  assumes "compact s"  "closed t"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5363
  shows "closed {x - y | x y. x \<in> s \<and> y \<in> t}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5364
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5365
  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
  5366
    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
  5367
  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
  5368
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5369
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5370
lemma closed_compact_differences:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5371
  fixes s t :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5372
  assumes "closed s" "compact t"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5373
  shows "closed {x - y | x y. x \<in> s \<and> y \<in> t}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5374
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5375
  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
  5376
    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
  5377
 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
  5378
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5379
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5380
lemma closed_translation:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5381
  fixes a :: "'a::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5382
  assumes "closed s"  shows "closed ((\<lambda>x. a + x) ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5383
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5384
  have "{a + y |y. y \<in> s} = (op + a ` s)" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5385
  thus ?thesis using compact_closed_sums[OF compact_sing[of a] assms] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5386
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5387
34105
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
  5388
lemma translation_Compl:
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
  5389
  fixes a :: "'a::ab_group_add"
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
  5390
  shows "(\<lambda>x. a + x) ` (- t) = - ((\<lambda>x. a + x) ` t)"
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
  5391
  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
  5392
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5393
lemma translation_UNIV:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5394
  fixes a :: "'a::ab_group_add" shows "range (\<lambda>x. a + x) = UNIV"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5395
  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
  5396
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5397
lemma translation_diff:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5398
  fixes a :: "'a::ab_group_add"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5399
  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
  5400
  by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5401
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5402
lemma closure_translation:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5403
  fixes a :: "'a::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5404
  shows "closure ((\<lambda>x. a + x) ` s) = (\<lambda>x. a + x) ` (closure s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5405
proof-
34105
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
  5406
  have *:"op + a ` (- s) = - op + a ` s"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5407
    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
  5408
  show ?thesis unfolding closure_interior translation_Compl
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
  5409
    using interior_translation[of a "- s"] unfolding * by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5410
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5411
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5412
lemma frontier_translation:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5413
  fixes a :: "'a::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5414
  shows "frontier((\<lambda>x. a + x) ` s) = (\<lambda>x. a + x) ` (frontier s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5415
  unfolding frontier_def translation_diff interior_translation closure_translation by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5416
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  5417
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  5418
subsection {* Separation between points and sets *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5419
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5420
lemma separate_point_closed:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5421
  fixes s :: "'a::heine_borel set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5422
  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
  5423
proof(cases "s = {}")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5424
  case True
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5425
  thus ?thesis by(auto intro!: exI[where x=1])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5426
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5427
  case False
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5428
  assume "closed s" "a \<notin> s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5429
  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
  5430
  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
  5431
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5432
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5433
lemma separate_compact_closed:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5434
  fixes s t :: "'a::{heine_borel, real_normed_vector} set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5435
    (* TODO: does this generalize to heine_borel? *)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5436
  assumes "compact s" and "closed t" and "s \<inter> t = {}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5437
  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
  5438
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5439
  have "0 \<notin> {x - y |x y. x \<in> s \<and> y \<in> t}" using assms(3) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5440
  then obtain d where "d>0" and d:"\<forall>x\<in>{x - y |x y. x \<in> s \<and> y \<in> t}. d \<le> dist 0 x"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5441
    using separate_point_closed[OF compact_closed_differences[OF assms(1,2)], of 0] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5442
  { fix x y assume "x\<in>s" "y\<in>t"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5443
    hence "x - y \<in> {x - y |x y. x \<in> s \<and> y \<in> t}" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5444
    hence "d \<le> dist (x - y) 0" using d[THEN bspec[where x="x - y"]] using dist_commute
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5445
      by (auto  simp add: dist_commute)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5446
    hence "d \<le> dist x y" unfolding dist_norm by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5447
  thus ?thesis using `d>0` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5448
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5449
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5450
lemma separate_closed_compact:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5451
  fixes s t :: "'a::{heine_borel, real_normed_vector} set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5452
  assumes "closed s" and "compact t" and "s \<inter> t = {}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5453
  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
  5454
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5455
  have *:"t \<inter> s = {}" using assms(3) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5456
  show ?thesis using separate_compact_closed[OF assms(2,1) *]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5457
    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
  5458
    by (auto simp add: dist_commute)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5459
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5460
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  5461
36439
a65320184de9 move intervals section heading
huffman
parents: 36438
diff changeset
  5462
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
  5463
  
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5464
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
  5465
  "{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
  5466
  "{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
  5467
  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
  5468
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5469
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
  5470
  "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
  5471
  "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
  5472
  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
  5473
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5474
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
  5475
 "({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
  5476
 "({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
  5477
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
  5478
  { 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
  5479
    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
  5480
    hence "a\<bullet>i < b\<bullet>i" by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5481
    hence False using as by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5482
  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
  5483
  { 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
  5484
    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
  5485
    { 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
  5486
      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
  5487
      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
  5488
        by (auto simp: inner_add_left) }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5489
    hence "{a <..< b} \<noteq> {}" using mem_interval(1)[of "?x" a b] by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5490
  ultimately show ?th1 by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5491
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
  5492
  { 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
  5493
    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
  5494
    hence "a\<bullet>i \<le> b\<bullet>i" by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5495
    hence False using as by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5496
  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
  5497
  { assume as:"\<forall>i\<in>Basis. \<not> (b\<bullet>i < a\<bullet>i)"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5498
    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
  5499
    { 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
  5500
      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
  5501
      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
  5502
        by (auto simp: inner_add_left) }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5503
    hence "{a .. b} \<noteq> {}" using mem_interval(2)[of "?x" a b] by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5504
  ultimately show ?th2 by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5505
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5506
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5507
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
  5508
  "{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
  5509
  "{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
  5510
  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
  5511
44584
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  5512
lemma interval_sing:
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  5513
  fixes a :: "'a::ordered_euclidean_space"
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  5514
  shows "{a .. a} = {a}" and "{a<..<a} = {}"
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  5515
  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
  5516
  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
  5517
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5518
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
  5519
 "(\<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
  5520
 "(\<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
  5521
 "(\<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
  5522
 "(\<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
  5523
  unfolding subset_eq[unfolded Ball_def] unfolding mem_interval
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  5524
  by (best intro: order_trans less_le_trans le_less_trans less_imp_le)+
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  5525
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  5526
lemma interval_open_subset_closed:
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  5527
  fixes a :: "'a::ordered_euclidean_space"
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  5528
  shows "{a<..<b} \<subseteq> {a .. b}"
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  5529
  unfolding subset_eq [unfolded Ball_def] mem_interval
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  5530
  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
  5531
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5532
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
  5533
 "{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
  5534
 "{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
  5535
 "{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
  5536
 "{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
  5537
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5538
  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
  5539
  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
  5540
  { 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
  5541
    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
  5542
    fix i :: 'a assume i:"i\<in>Basis"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5543
    (** 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
  5544
    { 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
  5545
      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
  5546
      { 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
  5547
        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
  5548
          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
  5549
          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
  5550
      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
  5551
      moreover
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5552
      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
  5553
        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
  5554
        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
  5555
        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
  5556
      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
  5557
    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
  5558
    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
  5559
    { 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
  5560
      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
  5561
      { 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
  5562
        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
  5563
          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
  5564
          by (auto simp add: as2) }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5565
      hence "?x\<in>{c<..<d}" unfolding mem_interval by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5566
      moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5567
      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
  5568
        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
  5569
        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
  5570
        by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5571
      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
  5572
    hence "b\<bullet>i \<ge> d\<bullet>i" by(rule ccontr)auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5573
    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
  5574
    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
  5575
  } 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
  5576
  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
  5577
    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
  5578
    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
  5579
    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
  5580
    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
  5581
    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
  5582
    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
  5583
    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
  5584
  { 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
  5585
    fix i :: 'a assume i:"i\<in>Basis"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5586
    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
  5587
    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
  5588
  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
  5589
    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
  5590
    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
  5591
qed
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5592
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5593
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
  5594
 "{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
  5595
  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
  5596
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
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
  5598
  "{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
  5599
  "{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
  5600
  "{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
  5601
  "{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
  5602
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
  5603
  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
  5604
  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
  5605
      (\<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
  5606
    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
  5607
  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
  5608
  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
  5609
  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
  5610
  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
  5611
  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
  5612
qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5613
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5614
(* Moved interval_open_subset_closed a bit upwards *)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5615
44250
9133bc634d9c simplify proofs of lemmas open_interval, closed_interval
huffman
parents: 44233
diff changeset
  5616
lemma open_interval[intro]:
9133bc634d9c simplify proofs of lemmas open_interval, closed_interval
huffman
parents: 44233
diff changeset
  5617
  fixes a b :: "'a::ordered_euclidean_space" shows "open {a<..<b}"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5618
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
  5619
  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
  5620
    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
  5621
      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
  5622
  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
  5623
    by (auto simp add: eucl_less [where 'a='a])
9133bc634d9c simplify proofs of lemmas open_interval, closed_interval
huffman
parents: 44233
diff changeset
  5624
  finally show "open {a<..<b}" .
9133bc634d9c simplify proofs of lemmas open_interval, closed_interval
huffman
parents: 44233
diff changeset
  5625
qed
9133bc634d9c simplify proofs of lemmas open_interval, closed_interval
huffman
parents: 44233
diff changeset
  5626
9133bc634d9c simplify proofs of lemmas open_interval, closed_interval
huffman
parents: 44233
diff changeset
  5627
lemma closed_interval[intro]:
9133bc634d9c simplify proofs of lemmas open_interval, closed_interval
huffman
parents: 44233
diff changeset
  5628
  fixes a b :: "'a::ordered_euclidean_space" shows "closed {a .. b}"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5629
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
  5630
  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
  5631
    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
  5632
      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
  5633
  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
  5634
    by (auto simp add: eucl_le [where 'a='a])
9133bc634d9c simplify proofs of lemmas open_interval, closed_interval
huffman
parents: 44233
diff changeset
  5635
  finally show "closed {a .. b}" .
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5636
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5637
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  5638
lemma interior_closed_interval [intro]:
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  5639
  fixes a b :: "'a::ordered_euclidean_space"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  5640
  shows "interior {a..b} = {a<..<b}" (is "?L = ?R")
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5641
proof(rule subset_antisym)
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  5642
  show "?R \<subseteq> ?L" using interval_open_subset_closed open_interval
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  5643
    by (rule interior_maximal)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5644
next
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  5645
  { fix x assume "x \<in> interior {a..b}"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  5646
    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
  5647
    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
  5648
    { 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
  5649
      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
  5650
           "dist (x + (e / 2) *\<^sub>R i) x < e"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5651
        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
  5652
        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
  5653
      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
  5654
                     "(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
  5655
        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
  5656
        and   e[THEN spec[where x="x + (e/2) *\<^sub>R i"]]
44584
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  5657
        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
  5658
      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
  5659
        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
  5660
    hence "x \<in> {a<..<b}" unfolding mem_interval by auto  }
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  5661
  thus "?L \<subseteq> ?R" ..
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5662
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5663
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5664
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
  5665
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
  5666
  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
  5667
  { 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
  5668
    { 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
  5669
      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
  5670
    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
  5671
    hence "norm x \<le> ?b" using norm_le_l1[of x] by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5672
  thus ?thesis unfolding interval and bounded_iff by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5673
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5674
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5675
lemma bounded_interval: fixes a :: "'a::ordered_euclidean_space" shows
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5676
 "bounded {a .. b} \<and> bounded {a<..<b}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5677
  using bounded_closed_interval[of a b]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5678
  using interval_open_subset_closed[of a b]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5679
  using bounded_subset[of "{a..b}" "{a<..<b}"]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5680
  by simp
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 not_interval_univ: fixes a :: "'a::ordered_euclidean_space" shows
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5683
 "({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
  5684
  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
  5685
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5686
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
  5687
  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
  5688
  by (auto simp: compact_eq_seq_compact_metric)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5689
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5690
lemma open_interval_midpoint: fixes a :: "'a::ordered_euclidean_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5691
  assumes "{a<..<b} \<noteq> {}" shows "((1/2) *\<^sub>R (a + b)) \<in> {a<..<b}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5692
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
  5693
  { 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
  5694
    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
  5695
      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
  5696
  thus ?thesis unfolding mem_interval by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5697
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5698
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5699
lemma open_closed_interval_convex: fixes x :: "'a::ordered_euclidean_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5700
  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
  5701
  shows "(e *\<^sub>R x + (1 - e) *\<^sub>R y) \<in> {a<..<b}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5702
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
  5703
  { 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
  5704
    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
  5705
    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
  5706
      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
  5707
      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
  5708
      using y unfolding mem_interval using i apply simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5709
      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
  5710
    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
  5711
    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
  5712
    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
  5713
    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
  5714
      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
  5715
      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
  5716
      using y unfolding mem_interval using i apply simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5717
      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
  5718
    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
  5719
    } 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
  5720
  thus ?thesis unfolding mem_interval by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5721
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5722
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5723
lemma closure_open_interval: fixes a :: "'a::ordered_euclidean_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5724
  assumes "{a<..<b} \<noteq> {}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5725
  shows "closure {a<..<b} = {a .. b}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5726
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
  5727
  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
  5728
  let ?c = "(1 / 2) *\<^sub>R (a + b)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5729
  { fix x assume as:"x \<in> {a .. b}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5730
    def f == "\<lambda>n::nat. x + (inverse (real n + 1)) *\<^sub>R (?c - x)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5731
    { 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
  5732
      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
  5733
      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
  5734
        x + (inverse (real n + 1)) *\<^sub>R (((1 / 2) *\<^sub>R (a + b)) - x)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5735
        by (auto simp add: algebra_simps)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5736
      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
  5737
      hence False using fn unfolding f_def using xc by auto  }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5738
    moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5739
    { assume "\<not> (f ---> x) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5740
      { fix e::real assume "e>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5741
        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
  5742
        then obtain N::nat where "inverse (real (N + 1)) < e" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5743
        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
  5744
        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
  5745
      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
  5746
        unfolding LIMSEQ_def by(auto simp add: dist_norm)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5747
      hence "(f ---> x) sequentially" unfolding f_def
44125
230a8665c919 mark some redundant theorems as legacy
huffman
parents: 44122
diff changeset
  5748
        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
  5749
        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
  5750
    ultimately have "x \<in> closure {a<..<b}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5751
      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
  5752
  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
  5753
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5754
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
lemma bounded_subset_open_interval_symmetric: fixes s::"('a::ordered_euclidean_space) set"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5756
  assumes "bounded s"  shows "\<exists>a. s \<subseteq> {-a<..<a}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5757
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5758
  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
  5759
  def a \<equiv> "(\<Sum>i\<in>Basis. (b + 1) *\<^sub>R i)::'a"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5760
  { 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
  5761
    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
  5762
    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
  5763
      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
  5764
  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
  5765
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5766
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5767
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
  5768
  fixes s :: "('a::ordered_euclidean_space) set"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5769
  shows "bounded s ==> (\<exists>a b. s \<subseteq> {a<..<b})"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5770
  by (auto dest!: bounded_subset_open_interval_symmetric)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5771
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5772
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
  5773
  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 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
  5777
  thus ?thesis using interval_open_subset_closed[of "-a" a] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5778
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5779
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5780
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
  5781
  fixes s :: "('a::ordered_euclidean_space) set"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5782
  shows "bounded s ==> (\<exists>a b. s \<subseteq> {a .. b})"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5783
  using bounded_subset_closed_interval_symmetric[of s] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5784
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5785
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
  5786
  fixes a b :: "'a::ordered_euclidean_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5787
  shows "frontier {a .. b} = {a .. b} - {a<..<b}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5788
  unfolding frontier_def unfolding interior_closed_interval and closure_closed[OF closed_interval] ..
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5789
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5790
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
  5791
  fixes a b :: "'a::ordered_euclidean_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5792
  shows "frontier {a<..<b} = (if {a<..<b} = {} then {} else {a .. b} - {a<..<b})"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5793
proof(cases "{a<..<b} = {}")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5794
  case True thus ?thesis using frontier_empty by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5795
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5796
  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
  5797
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5798
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
lemma inter_interval_mixed_eq_empty: fixes a :: "'a::ordered_euclidean_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5800
  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
  5801
  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
  5802
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5803
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5804
(* Some stuff for half-infinite intervals too; FIXME: notation?  *)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5805
37673
f69f4b079275 generalize more lemmas from ordered_euclidean_space to euclidean_space
huffman
parents: 37649
diff changeset
  5806
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
  5807
  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
  5808
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
  5809
  { 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
  5810
    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
  5811
    { 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
  5812
      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
  5813
        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
  5814
      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
  5815
        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
  5816
    hence "x\<bullet>i \<le> b\<bullet>i" by(rule ccontr)auto  }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5817
  thus ?thesis unfolding closed_limpt unfolding islimpt_approachable by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5818
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5819
37673
f69f4b079275 generalize more lemmas from ordered_euclidean_space to euclidean_space
huffman
parents: 37649
diff changeset
  5820
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
  5821
  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
  5822
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
  5823
  { 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
  5824
    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
  5825
    { 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
  5826
      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
  5827
        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
  5828
      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
  5829
    hence "a\<bullet>i \<le> x\<bullet>i" by(rule ccontr)auto  }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5830
  thus ?thesis unfolding closed_limpt unfolding islimpt_approachable by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5831
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5832
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
  5833
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
  5834
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
  5835
  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
  5836
    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
  5837
  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
  5838
    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
  5839
  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
  5840
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
  5841
50881
ae630bab13da renamed countable_basis_space to second_countable_topology
hoelzl
parents: 50526
diff changeset
  5842
instance euclidean_space \<subseteq> second_countable_topology
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  5843
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
  5844
  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
  5845
  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
  5846
  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
  5847
  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
  5848
  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
  5849
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5850
  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
  5851
    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
  5852
  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
  5853
  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
  5854
  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
  5855
  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
  5856
    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
  5857
    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
  5858
      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
  5859
      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
  5860
      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
  5861
      done
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  5862
  qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  5863
  ultimately
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
  5864
  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
  5865
qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  5866
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  5867
instance ordered_euclidean_space \<subseteq> polish_space ..
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  5868
36439
a65320184de9 move intervals section heading
huffman
parents: 36438
diff changeset
  5869
text {* Intervals in general, including infinite and mixtures of open and closed. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5870
37732
6432bf0d7191 generalize type of is_interval to class euclidean_space
huffman
parents: 37680
diff changeset
  5871
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
  5872
  (\<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
  5873
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5874
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
  5875
  "is_interval {a<..<b}" (is ?th2) proof -
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5876
  show ?th1 ?th2  unfolding is_interval_def mem_interval Ball_def atLeastAtMost_iff
44584
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  5877
    by(meson order_trans le_less_trans less_le_trans less_trans)+ qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5878
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5879
lemma is_interval_empty:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5880
 "is_interval {}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5881
  unfolding is_interval_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5882
  by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5883
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5884
lemma is_interval_univ:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5885
 "is_interval UNIV"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5886
  unfolding is_interval_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5887
  by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5888
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  5889
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  5890
subsection {* Closure of halfspaces and hyperplanes *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5891
44219
f738e3200e24 generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
huffman
parents: 44216
diff changeset
  5892
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
  5893
  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
  5894
proof -
f738e3200e24 generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
huffman
parents: 44216
diff changeset
  5895
  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
  5896
    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
  5897
  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
  5898
    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
  5899
  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
  5900
    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
  5901
qed
f738e3200e24 generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
huffman
parents: 44216
diff changeset
  5902
f738e3200e24 generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
huffman
parents: 44216
diff changeset
  5903
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
  5904
  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
  5905
  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
  5906
  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
  5907
44213
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5908
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
  5909
  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
  5910
  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
  5911
  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
  5912
  shows "open {x. f x < g x}"
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5913
proof -
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5914
  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
  5915
    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
  5916
    by (rule isCont_open_vimage)
44213
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5917
  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
  5918
    by auto
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5919
  finally show ?thesis .
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5920
qed
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5921
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5922
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
  5923
  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
  5924
  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
  5925
  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
  5926
  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
  5927
proof -
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5928
  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
  5929
    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
  5930
    by (rule isCont_closed_vimage)
44213
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5931
  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
  5932
    by auto
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5933
  finally show ?thesis .
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5934
qed
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5935
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5936
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
  5937
  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
  5938
  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
  5939
  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
  5940
  shows "closed {x. f x = g x}"
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5941
proof -
44216
903bfe95fece generalized lemma closed_Collect_eq
huffman
parents: 44213
diff changeset
  5942
  have "open {(x::'b, y::'b). x \<noteq> y}"
903bfe95fece generalized lemma closed_Collect_eq
huffman
parents: 44213
diff changeset
  5943
    unfolding open_prod_def by (auto dest!: hausdorff)
903bfe95fece generalized lemma closed_Collect_eq
huffman
parents: 44213
diff changeset
  5944
  hence "closed {(x::'b, y::'b). x = y}"
903bfe95fece generalized lemma closed_Collect_eq
huffman
parents: 44213
diff changeset
  5945
    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
  5946
  with isCont_Pair [OF f g]
44216
903bfe95fece generalized lemma closed_Collect_eq
huffman
parents: 44213
diff changeset
  5947
  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
  5948
    by (rule isCont_closed_vimage)
44216
903bfe95fece generalized lemma closed_Collect_eq
huffman
parents: 44213
diff changeset
  5949
  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
  5950
  finally show ?thesis .
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5951
qed
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5952
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5953
lemma continuous_at_inner: "continuous (at x) (inner a)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5954
  unfolding continuous_at by (intro tendsto_intros)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5955
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5956
lemma closed_halfspace_le: "closed {x. inner a x \<le> b}"
44233
aa74ce315bae add simp rules for isCont
huffman
parents: 44219
diff changeset
  5957
  by (simp add: closed_Collect_le)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5958
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5959
lemma closed_halfspace_ge: "closed {x. inner a x \<ge> b}"
44233
aa74ce315bae add simp rules for isCont
huffman
parents: 44219
diff changeset
  5960
  by (simp add: closed_Collect_le)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5961
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5962
lemma closed_hyperplane: "closed {x. inner a x = b}"
44233
aa74ce315bae add simp rules for isCont
huffman
parents: 44219
diff changeset
  5963
  by (simp add: closed_Collect_eq)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5964
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5965
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
  5966
  shows "closed {x::'a::euclidean_space. x\<bullet>i \<le> a}"
44233
aa74ce315bae add simp rules for isCont
huffman
parents: 44219
diff changeset
  5967
  by (simp add: closed_Collect_le)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5968
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5969
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
  5970
  shows "closed {x::'a::euclidean_space. x\<bullet>i \<ge> a}"
44233
aa74ce315bae add simp rules for isCont
huffman
parents: 44219
diff changeset
  5971
  by (simp add: closed_Collect_le)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5972
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  5973
text {* Openness of halfspaces. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5974
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5975
lemma open_halfspace_lt: "open {x. inner a x < b}"
44233
aa74ce315bae add simp rules for isCont
huffman
parents: 44219
diff changeset
  5976
  by (simp add: open_Collect_less)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5977
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5978
lemma open_halfspace_gt: "open {x. inner a x > b}"
44233
aa74ce315bae add simp rules for isCont
huffman
parents: 44219
diff changeset
  5979
  by (simp add: open_Collect_less)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5980
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5981
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
  5982
  shows "open {x::'a::euclidean_space. x\<bullet>i < a}"
44233
aa74ce315bae add simp rules for isCont
huffman
parents: 44219
diff changeset
  5983
  by (simp add: open_Collect_less)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5984
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5985
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
  5986
  shows "open {x::'a::euclidean_space. x\<bullet>i > a}"
44233
aa74ce315bae add simp rules for isCont
huffman
parents: 44219
diff changeset
  5987
  by (simp add: open_Collect_less)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5988
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5989
text{* Instantiation for intervals on @{text ordered_euclidean_space} *}
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5990
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5991
lemma eucl_lessThan_eq_halfspaces:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5992
  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
  5993
  shows "{..<a} = (\<Inter>i\<in>Basis. {x. x \<bullet> i < a \<bullet> i})"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5994
 by (auto simp: eucl_less[where 'a='a])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5995
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5996
lemma eucl_greaterThan_eq_halfspaces:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5997
  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
  5998
  shows "{a<..} = (\<Inter>i\<in>Basis. {x. a \<bullet> i < x \<bullet> i})"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5999
 by (auto simp: eucl_less[where 'a='a])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  6000
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  6001
lemma eucl_atMost_eq_halfspaces:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  6002
  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
  6003
  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
  6004
 by (auto simp: eucl_le[where 'a='a])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  6005
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  6006
lemma eucl_atLeast_eq_halfspaces:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  6007
  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
  6008
  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
  6009
 by (auto simp: eucl_le[where 'a='a])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  6010
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  6011
lemma open_eucl_lessThan[simp, intro]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  6012
  fixes a :: "'a\<Colon>ordered_euclidean_space"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  6013
  shows "open {..< a}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  6014
  by (auto simp: eucl_lessThan_eq_halfspaces open_halfspace_component_lt)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  6015
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  6016
lemma open_eucl_greaterThan[simp, intro]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  6017
  fixes a :: "'a\<Colon>ordered_euclidean_space"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  6018
  shows "open {a <..}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  6019
  by (auto simp: eucl_greaterThan_eq_halfspaces open_halfspace_component_gt)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  6020
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  6021
lemma closed_eucl_atMost[simp, intro]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  6022
  fixes a :: "'a\<Colon>ordered_euclidean_space"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  6023
  shows "closed {.. a}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  6024
  unfolding eucl_atMost_eq_halfspaces
44233
aa74ce315bae add simp rules for isCont
huffman
parents: 44219
diff changeset
  6025
  by (simp add: closed_INT closed_Collect_le)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  6026
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  6027
lemma closed_eucl_atLeast[simp, intro]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  6028
  fixes a :: "'a\<Colon>ordered_euclidean_space"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  6029
  shows "closed {a ..}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  6030
  unfolding eucl_atLeast_eq_halfspaces
44233
aa74ce315bae add simp rules for isCont
huffman
parents: 44219
diff changeset
  6031
  by (simp add: closed_INT closed_Collect_le)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  6032
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  6033
text {* This gives a simple derivation of limit component bounds. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6034
37673
f69f4b079275 generalize more lemmas from ordered_euclidean_space to euclidean_space
huffman
parents: 37649
diff changeset
  6035
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
  6036
  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
  6037
  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
  6038
  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
  6039
37673
f69f4b079275 generalize more lemmas from ordered_euclidean_space to euclidean_space
huffman
parents: 37649
diff changeset
  6040
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
  6041
  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
  6042
  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
  6043
  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
  6044
37673
f69f4b079275 generalize more lemmas from ordered_euclidean_space to euclidean_space
huffman
parents: 37649
diff changeset
  6045
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
  6046
  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
  6047
  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
  6048
  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
  6049
  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
  6050
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6051
text{* Limits relative to a union.                                               *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6052
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6053
lemma eventually_within_Un:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6054
  "eventually P (net within (s \<union> t)) \<longleftrightarrow>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6055
    eventually P (net within s) \<and> eventually P (net within t)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6056
  unfolding Limits.eventually_within
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6057
  by (auto elim!: eventually_rev_mp)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6058
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6059
lemma Lim_within_union:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6060
 "(f ---> l) (net within (s \<union> t)) \<longleftrightarrow>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6061
  (f ---> l) (net within s) \<and> (f ---> l) (net within t)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6062
  unfolding tendsto_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6063
  by (auto simp add: eventually_within_Un)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6064
36442
b96d9dc6acca generalize more continuity lemmas
huffman
parents: 36441
diff changeset
  6065
lemma Lim_topological:
b96d9dc6acca generalize more continuity lemmas
huffman
parents: 36441
diff changeset
  6066
 "(f ---> l) net \<longleftrightarrow>
b96d9dc6acca generalize more continuity lemmas
huffman
parents: 36441
diff changeset
  6067
        trivial_limit net \<or>
b96d9dc6acca generalize more continuity lemmas
huffman
parents: 36441
diff changeset
  6068
        (\<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
  6069
  unfolding tendsto_def trivial_limit_eq by auto
b96d9dc6acca generalize more continuity lemmas
huffman
parents: 36441
diff changeset
  6070
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6071
lemma continuous_on_union:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6072
  assumes "closed s" "closed t" "continuous_on s f" "continuous_on t f"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6073
  shows "continuous_on (s \<union> t) f"
36442
b96d9dc6acca generalize more continuity lemmas
huffman
parents: 36441
diff changeset
  6074
  using assms unfolding continuous_on Lim_within_union
b96d9dc6acca generalize more continuity lemmas
huffman
parents: 36441
diff changeset
  6075
  unfolding Lim_topological trivial_limit_within closed_limpt by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6076
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6077
lemma continuous_on_cases:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6078
  assumes "closed s" "closed t" "continuous_on s f" "continuous_on t g"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6079
          "\<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
  6080
  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
  6081
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6082
  let ?h = "(\<lambda>x. if P x then f x else g x)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6083
  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
  6084
  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
  6085
  moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6086
  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
  6087
  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
  6088
  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
  6089
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6090
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6091
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6092
text{* Some more convenient intermediate-value theorem formulations.             *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6093
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6094
lemma connected_ivt_hyperplane:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6095
  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
  6096
  shows "\<exists>z \<in> s. inner a z = b"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6097
proof(rule ccontr)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6098
  assume as:"\<not> (\<exists>z\<in>s. inner a z = b)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6099
  let ?A = "{x. inner a x < b}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6100
  let ?B = "{x. inner a x > b}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6101
  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
  6102
  moreover have "?A \<inter> ?B = {}" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6103
  moreover have "s \<subseteq> ?A \<union> ?B" using as by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6104
  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
  6105
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6106
37673
f69f4b079275 generalize more lemmas from ordered_euclidean_space to euclidean_space
huffman
parents: 37649
diff changeset
  6107
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
  6108
 "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
  6109
  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
  6110
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  6111
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  6112
subsection {* Homeomorphisms *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6113
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6114
definition "homeomorphism s t f g \<equiv>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6115
     (\<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
  6116
     (\<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
  6117
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6118
definition
50898
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  6119
  homeomorphic :: "'a::topological_space set \<Rightarrow> 'b::topological_space set \<Rightarrow> bool"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6120
    (infixr "homeomorphic" 60) where
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6121
  homeomorphic_def: "s homeomorphic t \<equiv> (\<exists>f g. homeomorphism s t f g)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6122
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6123
lemma homeomorphic_refl: "s homeomorphic s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6124
  unfolding homeomorphic_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6125
  unfolding homeomorphism_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6126
  using continuous_on_id
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6127
  apply(rule_tac x = "(\<lambda>x. x)" in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6128
  apply(rule_tac x = "(\<lambda>x. x)" in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6129
  by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6130
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6131
lemma homeomorphic_sym:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6132
 "s homeomorphic t \<longleftrightarrow> t homeomorphic s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6133
unfolding homeomorphic_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6134
unfolding homeomorphism_def
33324
51eb2ffa2189 Tidied up some very ugly proofs
paulson
parents: 33270
diff changeset
  6135
by blast 
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6136
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6137
lemma homeomorphic_trans:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6138
  assumes "s homeomorphic t" "t homeomorphic u" shows "s homeomorphic u"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6139
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6140
  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
  6141
    using assms(1) unfolding homeomorphic_def homeomorphism_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6142
  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
  6143
    using assms(2) unfolding homeomorphic_def homeomorphism_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6144
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6145
  { 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
  6146
  moreover have "(f2 \<circ> f1) ` s = u" using fg1(2) fg2(2) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6147
  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
  6148
  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
  6149
  moreover have "(g1 \<circ> g2) ` u = s" using fg1(5) fg2(5) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6150
  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
  6151
  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
  6152
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6153
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6154
lemma homeomorphic_minimal:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6155
 "s homeomorphic t \<longleftrightarrow>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6156
    (\<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
  6157
           (\<forall>y\<in>t. g(y) \<in> s \<and> (f(g(y)) = y)) \<and>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6158
           continuous_on s f \<and> continuous_on t g)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6159
unfolding homeomorphic_def homeomorphism_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6160
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
  6161
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
  6162
unfolding image_iff
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6163
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
  6164
apply auto apply(rule_tac x="g x" in bexI) apply auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6165
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
  6166
apply auto apply(rule_tac x="f x" in bexI) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6167
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  6168
text {* Relatively weak hypotheses if a set is compact. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6169
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6170
lemma homeomorphism_compact:
50898
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  6171
  fixes f :: "'a::topological_space \<Rightarrow> 'b::t2_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6172
  assumes "compact s" "continuous_on s f"  "f ` s = t"  "inj_on f s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6173
  shows "\<exists>g. homeomorphism s t f g"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6174
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6175
  def g \<equiv> "\<lambda>x. SOME y. y\<in>s \<and> f y = x"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6176
  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
  6177
  { fix y assume "y\<in>t"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6178
    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
  6179
    hence "g (f x) = x" using g by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6180
    hence "f (g y) = y" unfolding x(1)[THEN sym] by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6181
  hence g':"\<forall>x\<in>t. f (g x) = x" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6182
  moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6183
  { fix x
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6184
    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
  6185
    moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6186
    { assume "x\<in>g ` t"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6187
      then obtain y where y:"y\<in>t" "g y = x" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6188
      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
  6189
      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
  6190
    ultimately have "x\<in>s \<longleftrightarrow> x \<in> g ` t" ..  }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6191
  hence "g ` t = s" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6192
  ultimately
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6193
  show ?thesis unfolding homeomorphism_def homeomorphic_def
44647
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  6194
    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
  6195
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6196
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6197
lemma homeomorphic_compact:
50898
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  6198
  fixes f :: "'a::topological_space \<Rightarrow> 'b::t2_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6199
  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
  6200
          \<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
  6201
  unfolding homeomorphic_def by (metis homeomorphism_compact)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6202
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6203
text{* Preservation of topological properties.                                   *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6204
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6205
lemma homeomorphic_compactness:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6206
 "s homeomorphic t ==> (compact s \<longleftrightarrow> compact t)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6207
unfolding homeomorphic_def homeomorphism_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6208
by (metis compact_continuous_image)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6209
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6210
text{* Results on translation, scaling etc.                                      *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6211
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6212
lemma homeomorphic_scaling:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6213
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6214
  assumes "c \<noteq> 0"  shows "s homeomorphic ((\<lambda>x. c *\<^sub>R x) ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6215
  unfolding homeomorphic_minimal
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6216
  apply(rule_tac x="\<lambda>x. c *\<^sub>R x" in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6217
  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
  6218
  using assms by (auto simp add: continuous_on_intros)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6219
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6220
lemma homeomorphic_translation:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6221
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6222
  shows "s homeomorphic ((\<lambda>x. a + x) ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6223
  unfolding homeomorphic_minimal
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6224
  apply(rule_tac x="\<lambda>x. a + x" in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6225
  apply(rule_tac x="\<lambda>x. -a + x" in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6226
  using continuous_on_add[OF continuous_on_const continuous_on_id] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6227
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6228
lemma homeomorphic_affinity:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6229
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6230
  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
  6231
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6232
  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
  6233
  show ?thesis
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6234
    using homeomorphic_trans
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6235
    using homeomorphic_scaling[OF assms, of s]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6236
    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
  6237
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6238
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6239
lemma homeomorphic_balls:
50898
ebd9b82537e0 generalized more topology theorems
huffman
parents: 50897
diff changeset
  6240
  fixes a b ::"'a::real_normed_vector"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6241
  assumes "0 < d"  "0 < e"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6242
  shows "(ball a d) homeomorphic  (ball b e)" (is ?th)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6243
        "(cball a d) homeomorphic (cball b e)" (is ?cth)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6244
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6245
  show ?th unfolding homeomorphic_minimal
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6246
    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
  6247
    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
  6248
    using assms apply (auto simp add: dist_commute)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6249
    unfolding dist_norm
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6250
    apply (auto simp add: pos_divide_less_eq mult_strict_left_mono)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6251
    unfolding continuous_on
36659
f794e92784aa adapt to removed premise on tendsto lemma (cf. 88f0125c3bd2)
huffman
parents: 36623
diff changeset
  6252
    by (intro ballI tendsto_intros, simp)+
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6253
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6254
  show ?cth unfolding homeomorphic_minimal
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6255
    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
  6256
    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
  6257
    using assms apply (auto simp add: dist_commute)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6258
    unfolding dist_norm
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6259
    apply (auto simp add: pos_divide_le_eq)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6260
    unfolding continuous_on
36659
f794e92784aa adapt to removed premise on tendsto lemma (cf. 88f0125c3bd2)
huffman
parents: 36623
diff changeset
  6261
    by (intro ballI tendsto_intros, simp)+
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6262
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6263
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6264
text{* "Isometry" (up to constant bounds) of injective linear map etc.           *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6265
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6266
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
  6267
  fixes x :: "nat \<Rightarrow> 'a::euclidean_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6268
  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
  6269
  shows "Cauchy x"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6270
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6271
  interpret f: bounded_linear f by fact
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6272
  { fix d::real assume "d>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6273
    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
  6274
      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
  6275
    { 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
  6276
      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
  6277
        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
  6278
        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
  6279
      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
  6280
        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
  6281
      finally have "norm (x n - x N) < d" using `e>0` by simp }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6282
    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
  6283
  thus ?thesis unfolding cauchy and dist_norm by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6284
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6285
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6286
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
  6287
  fixes f :: "'a::euclidean_space => 'b::euclidean_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6288
  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
  6289
  shows "complete(f ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6290
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6291
  { 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
  6292
    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
  6293
      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
  6294
    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
  6295
    hence "f \<circ> x = g" unfolding fun_eq_iff by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6296
    then obtain l where "l\<in>s" and l:"(x ---> l) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6297
      using cs[unfolded complete_def, THEN spec[where x="x"]]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6298
      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
  6299
    hence "\<exists>l\<in>f ` s. (g ---> l) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6300
      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
  6301
      unfolding `f \<circ> x = g` by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6302
  thus ?thesis unfolding complete_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6303
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6304
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6305
lemma dist_0_norm:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6306
  fixes x :: "'a::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6307
  shows "dist 0 x = norm x"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6308
unfolding dist_norm by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6309
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  6310
lemma injective_imp_isometric: fixes f::"'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6311
  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
  6312
  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
  6313
proof(cases "s \<subseteq> {0::'a}")
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6314
  case True
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6315
  { fix x assume "x \<in> s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6316
    hence "x = 0" using True by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6317
    hence "norm x \<le> norm (f x)" by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6318
  thus ?thesis by(auto intro!: exI[where x=1])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6319
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6320
  interpret f: bounded_linear f by fact
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6321
  case False
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6322
  then obtain a where a:"a\<noteq>0" "a\<in>s" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6323
  from False have "s \<noteq> {}" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6324
  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
  6325
  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
  6326
  let ?S'' = "{x::'a. norm x = norm a}"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6327
36362
06475a1547cb fix lots of looping simp calls and other warnings
huffman
parents: 36360
diff changeset
  6328
  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
  6329
  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
  6330
  moreover have "?S' = s \<inter> ?S''" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6331
  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
  6332
  moreover have *:"f ` ?S' = ?S" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6333
  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
  6334
  hence "closed ?S" using compact_imp_closed by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6335
  moreover have "?S \<noteq> {}" using a by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6336
  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
  6337
  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
  6338
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6339
  let ?e = "norm (f b) / norm b"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6340
  have "norm b > 0" using ba and a and norm_ge_zero by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6341
  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
  6342
  ultimately have "0 < norm (f b) / norm b" by(simp only: divide_pos_pos)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6343
  moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6344
  { fix x assume "x\<in>s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6345
    hence "norm (f b) / norm b * norm x \<le> norm (f x)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6346
    proof(cases "x=0")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6347
      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
  6348
    next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6349
      case False
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6350
      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
  6351
      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
  6352
      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
  6353
      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
  6354
        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
  6355
        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
  6356
    qed }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6357
  ultimately
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6358
  show ?thesis by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6359
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6360
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6361
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
  6362
  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6363
  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
  6364
  shows "closed(f ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6365
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6366
  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
  6367
  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
  6368
    unfolding complete_eq_closed[THEN sym] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6369
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6370
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  6371
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  6372
subsection {* Some properties of a canonical subspace *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6373
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6374
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
  6375
  "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
  6376
  unfolding subspace_def by (auto simp: inner_add_left)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6377
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6378
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
  6379
 "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
  6380
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
  6381
  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
  6382
  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
  6383
    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
  6384
  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
  6385
    by auto
d366fa5551ef declare euclidean_simps [simp] at the point they are proved;
huffman
parents: 44365
diff changeset
  6386
  finally show "closed ?A" .
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6387
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6388
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
  6389
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
  6390
  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
  6391
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
  6392
  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
  6393
  have "d \<subseteq> ?A" using d by (auto simp: inner_Basis)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6394
  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
  6395
  { 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
  6396
    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
  6397
    from this d have "x \<in> span d"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6398
    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
  6399
      case empty hence "x=0" apply(rule_tac euclidean_eqI) by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6400
      thus ?case using subspace_0[OF subspace_span[of "{}"]] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6401
    next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6402
      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
  6403
      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
  6404
      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
  6405
      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
  6406
      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
  6407
      { 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
  6408
        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
  6409
          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
  6410
      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
  6411
      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
  6412
        using span_mono[of F "insert k F"] using assms by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6413
      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
  6414
      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
  6415
      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
  6416
        using span_mul by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6417
      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
  6418
      have "y + (x\<bullet>k) *\<^sub>R k \<in> span (insert k F)"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6419
        using span_add by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6420
      thus ?case using y by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6421
    qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6422
  }
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
  hence "?A \<subseteq> span d" by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6424
  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
  6425
  { 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
  6426
  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
  6427
  moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6428
  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
  6429
  ultimately show ?thesis using dim_unique[of d ?A] by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6430
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6431
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6432
text{* Hence closure and completeness of all subspaces.                          *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6433
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
  6434
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
  6435
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
  6436
  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
  6437
  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
  6438
  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
  6439
    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
  6440
  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
  6441
    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
  6442
  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
  6443
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
  6444
  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
  6445
qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6446
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  6447
lemma closed_subspace: fixes s::"('a::euclidean_space) set"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6448
  assumes "subspace s" shows "closed s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6449
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
  6450
  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
  6451
  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
  6452
  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
  6453
  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
  6454
      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
  6455
    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
  6456
    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
  6457
  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
  6458
  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
  6459
  { 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
  6460
  moreover have "closed ?t" using closed_substandard .
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6461
  moreover have "subspace ?t" using subspace_substandard .
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6462
  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
  6463
    unfolding f(2) using f(1) unfolding linear_conv_bounded_linear by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6464
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6465
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6466
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
  6467
  fixes s :: "('a::euclidean_space) set" shows "subspace s ==> complete s"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6468
  using complete_eq_closed closed_subspace
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6469
  by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6470
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6471
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
  6472
  fixes s :: "('a::euclidean_space) set"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6473
  shows "dim(closure s) = dim s" (is "?dc = ?d")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6474
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6475
  have "?dc \<le> ?d" using closure_minimal[OF span_inc, of s]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6476
    using closed_subspace[OF subspace_span, of s]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6477
    using dim_subset[of "closure s" "span s"] unfolding dim_span by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6478
  thus ?thesis using dim_subset[OF closure_subset, of s] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6479
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6480
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  6481
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  6482
subsection {* Affine transformations of intervals *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6483
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6484
lemma real_affinity_le:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34999
diff changeset
  6485
 "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
  6486
  by (simp add: field_simps inverse_eq_divide)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6487
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6488
lemma real_le_affinity:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34999
diff changeset
  6489
 "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
  6490
  by (simp add: field_simps inverse_eq_divide)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6491
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6492
lemma real_affinity_lt:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34999
diff changeset
  6493
 "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
  6494
  by (simp add: field_simps inverse_eq_divide)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6495
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6496
lemma real_lt_affinity:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34999
diff changeset
  6497
 "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
  6498
  by (simp add: field_simps inverse_eq_divide)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6499
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6500
lemma real_affinity_eq:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34999
diff changeset
  6501
 "(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
  6502
  by (simp add: field_simps inverse_eq_divide)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6503
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6504
lemma real_eq_affinity:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34999
diff changeset
  6505
 "(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
  6506
  by (simp add: field_simps inverse_eq_divide)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6507
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6508
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
  6509
  fixes a b c :: "'a::ordered_euclidean_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6510
  shows "(\<lambda>x. m *\<^sub>R x + c) ` {a .. b} =
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6511
            (if {a .. b} = {} then {}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6512
            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
  6513
            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
  6514
proof(cases "m=0")  
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6515
  { 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
  6516
    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
  6517
      apply(subst euclidean_eq_iff) by (auto intro: order_antisym) }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6518
  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
  6519
  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
  6520
  ultimately show ?thesis by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6521
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6522
  case False
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6523
  { fix y assume "a \<le> y" "y \<le> b" "m > 0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6524
    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
  6525
      unfolding eucl_le[where 'a='a] by (auto simp: inner_simps)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6526
  } moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6527
  { fix y assume "a \<le> y" "y \<le> b" "m < 0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6528
    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
  6529
      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
  6530
  } moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6531
  { 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
  6532
    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
  6533
      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
  6534
      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
  6535
      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
  6536
  } moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6537
  { 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
  6538
    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
  6539
      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
  6540
      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
  6541
      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
  6542
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6543
  ultimately show ?thesis using False by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6544
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6545
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  6546
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
  6547
  (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
  6548
  using image_affinity_interval[of m 0 a b] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6549
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  6550
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  6551
subsection {* Banach fixed point theorem (not really topological...) *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6552
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6553
lemma banach_fix:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6554
  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
  6555
          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
  6556
  shows "\<exists>! x\<in>s. (f x = x)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6557
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6558
  have "1 - c > 0" using c by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6559
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6560
  from s(2) obtain z0 where "z0 \<in> s" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6561
  def z \<equiv> "\<lambda>n. (f ^^ n) z0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6562
  { fix n::nat
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6563
    have "z n \<in> s" unfolding z_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6564
    proof(induct n) case 0 thus ?case using `z0 \<in>s` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6565
    next case Suc thus ?case using f by auto qed }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6566
  note z_in_s = this
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6567
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6568
  def d \<equiv> "dist (z 0) (z 1)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6569
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6570
  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
  6571
  { fix n::nat
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6572
    have "dist (z n) (z (Suc n)) \<le> (c ^ n) * d"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6573
    proof(induct n)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6574
      case 0 thus ?case unfolding d_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6575
    next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6576
      case (Suc m)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6577
      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
  6578
        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
  6579
      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
  6580
        unfolding fzn and mult_le_cancel_left by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6581
    qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6582
  } note cf_z = this
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6583
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6584
  { fix n m::nat
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6585
    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
  6586
    proof(induct n)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6587
      case 0 show ?case by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6588
    next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6589
      case (Suc k)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6590
      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
  6591
        using dist_triangle and c by(auto simp add: dist_triangle)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6592
      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
  6593
        using cf_z[of "m + k"] and c by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6594
      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
  6595
        using Suc by (auto simp add: field_simps)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6596
      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
  6597
        unfolding power_add by (auto simp add: field_simps)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6598
      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
  6599
        using c by (auto simp add: field_simps)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6600
      finally show ?case by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6601
    qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6602
  } note cf_z2 = this
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6603
  { fix e::real assume "e>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6604
    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
  6605
    proof(cases "d = 0")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6606
      case True
41863
e5104b436ea1 removed dependency on Dense_Linear_Order
boehmes
parents: 41413
diff changeset
  6607
      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
  6608
        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
  6609
      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
  6610
        by (simp add: *)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6611
      thus ?thesis using `e>0` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6612
    next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6613
      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
  6614
        by (metis False d_def less_le)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6615
      hence "0 < e * (1 - c) / d" using `e>0` and `1-c>0`
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6616
        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
  6617
      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
  6618
      { 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
  6619
        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
  6620
        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
  6621
        hence **:"d * (1 - c ^ (m - n)) / (1 - c) > 0"
36778
739a9379e29b avoid using real-specific versions of generic lemmas
huffman
parents: 36669
diff changeset
  6622
          using mult_pos_pos[OF `d>0`, of "1 - c ^ (m - n)"]
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6623
          using divide_pos_pos[of "d * (1 - c ^ (m - n))" "1 - c"]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6624
          using `0 < 1 - c` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6625
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6626
        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
  6627
          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
  6628
          by (auto simp add: mult_commute dist_commute)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6629
        also have "\<dots> \<le> c ^ N * d * (1 - c ^ (m - n)) / (1 - c)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6630
          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
  6631
          unfolding mult_assoc by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6632
        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
  6633
          using mult_strict_right_mono[OF N **] unfolding mult_assoc by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6634
        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
  6635
        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
  6636
        finally have  "dist (z m) (z n) < e" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6637
      } note * = this
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6638
      { fix m n::nat assume as:"N\<le>m" "N\<le>n"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6639
        hence "dist (z n) (z m) < e"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6640
        proof(cases "n = m")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6641
          case True thus ?thesis using `e>0` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6642
        next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6643
          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
  6644
        qed }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6645
      thus ?thesis by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6646
    qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6647
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6648
  hence "Cauchy z" unfolding cauchy_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6649
  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
  6650
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6651
  def e \<equiv> "dist (f x) x"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6652
  have "e = 0" proof(rule ccontr)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6653
    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
  6654
      by (metis dist_eq_0_iff dist_nz e_def)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6655
    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
  6656
      using x[unfolded LIMSEQ_def, THEN spec[where x="e/2"]] by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6657
    hence N':"dist (z N) x < e / 2" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6658
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6659
    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
  6660
      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
  6661
      by (metis dist_eq_0_iff dist_nz order_less_asym less_le)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6662
    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
  6663
      using z_in_s[of N] `x\<in>s` using c by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6664
    also have "\<dots> < e / 2" using N' and c using * by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6665
    finally show False unfolding fzn
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6666
      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
  6667
      unfolding e_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6668
  qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6669
  hence "f x = x" unfolding e_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6670
  moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6671
  { fix y assume "f y = y" "y\<in>s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6672
    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
  6673
      using `x\<in>s` and `f x = x` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6674
    hence "dist x y = 0" unfolding mult_le_cancel_right1
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6675
      using c and zero_le_dist[of x y] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6676
    hence "y = x" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6677
  }
34999
5312d2ffee3b Changed 'bounded unique existential quantifiers' from a constant to syntax translation.
hoelzl
parents: 34964
diff changeset
  6678
  ultimately show ?thesis using `x\<in>s` by blast+
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6679
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6680
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  6681
subsection {* Edelstein fixed point theorem *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6682
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6683
lemma edelstein_fix:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6684
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6685
  assumes s:"compact s" "s \<noteq> {}" and gs:"(g ` s) \<subseteq> s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6686
      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
  6687
  shows "\<exists>! x\<in>s. g x = x"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6688
proof(cases "\<exists>x\<in>s. g x \<noteq> x")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6689
  obtain x where "x\<in>s" using s(2) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6690
  case False hence g:"\<forall>x\<in>s. g x = x" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6691
  { fix y assume "y\<in>s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6692
    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
  6693
      unfolding g[THEN bspec[where x=x], OF `x\<in>s`]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6694
      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
  6695
  thus ?thesis using `x\<in>s` and g by blast+
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6696
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6697
  case True
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6698
  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
  6699
  { fix x y assume "x \<in> s" "y \<in> s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6700
    hence "dist (g x) (g y) \<le> dist x y"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6701
      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
  6702
  def y \<equiv> "g x"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6703
  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
  6704
  def f \<equiv> "\<lambda>n. g ^^ n"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6705
  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
  6706
  have [simp]:"\<And>z. f 0 z = z" unfolding f_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6707
  { fix n::nat and z assume "z\<in>s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6708
    have "f n z \<in> s" unfolding f_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6709
    proof(induct n)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6710
      case 0 thus ?case using `z\<in>s` by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6711
    next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6712
      case (Suc n) thus ?case using gs[unfolded image_subset_iff] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6713
    qed } note fs = this
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6714
  { fix m n ::nat assume "m\<le>n"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6715
    fix w z assume "w\<in>s" "z\<in>s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6716
    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
  6717
    proof(induct n)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6718
      case 0 thus ?case by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6719
    next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6720
      case (Suc n)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6721
      thus ?case proof(cases "m\<le>n")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6722
        case True thus ?thesis using Suc(1)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6723
          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
  6724
      next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6725
        case False hence mn:"m = Suc n" using Suc(2) by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6726
        show ?thesis unfolding mn  by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6727
      qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6728
    qed } note distf = this
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6729
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6730
  def h \<equiv> "\<lambda>n. (f n x, f n y)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6731
  let ?s2 = "s \<times> s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6732
  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
  6733
    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
  6734
    using fs[OF `x\<in>s`] and fs[OF `y\<in>s`] by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6735
  def a \<equiv> "fst l" def b \<equiv> "snd l"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6736
  have lab:"l = (a, b)" unfolding a_def b_def by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6737
  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
  6738
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6739
  have lima:"((fst \<circ> (h \<circ> r)) ---> a) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6740
   and limb:"((snd \<circ> (h \<circ> r)) ---> b) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6741
    using lr
44167
e81d676d598e avoid duplicate rule warnings
huffman
parents: 44139
diff changeset
  6742
    unfolding o_def a_def b_def by (rule tendsto_intros)+
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6743
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6744
  { fix n::nat
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6745
    have *:"\<And>fx fy (x::'a) y. dist fx fy \<le> dist x y \<Longrightarrow> \<not> (dist (fx - fy) (a - b) < dist a b - dist x y)" unfolding dist_norm by norm
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6746
    { fix x y :: 'a
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6747
      have "dist (-x) (-y) = dist x y" unfolding dist_norm
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6748
        using norm_minus_cancel[of "x - y"] by (auto simp add: uminus_add_conv_diff) } note ** = this
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6749
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6750
    { assume as:"dist a b > dist (f n x) (f n y)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6751
      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
  6752
        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
  6753
        using lima limb unfolding h_def LIMSEQ_def by (fastforce simp del: less_divide_eq_numeral1)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6754
      hence "dist (f (r (Na + Nb + n)) x - f (r (Na + Nb + n)) y) (a - b) < dist a b - dist (f n x) (f n y)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6755
        apply(erule_tac x="Na+Nb+n" in allE)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6756
        apply(erule_tac x="Na+Nb+n" in allE) apply simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6757
        using dist_triangle_add_half[of a "f (r (Na + Nb + n)) x" "dist a b - dist (f n x) (f n y)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6758
          "-b"  "- f (r (Na + Nb + n)) y"]
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36336
diff changeset
  6759
        unfolding ** by (auto simp add: algebra_simps dist_commute)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6760
      moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6761
      have "dist (f (r (Na + Nb + n)) x - f (r (Na + Nb + n)) y) (a - b) \<ge> dist a b - dist (f n x) (f n y)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6762
        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
  6763
        using seq_suble[OF r, of "Na+Nb+n"]
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6764
        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
  6765
      ultimately have False by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6766
    }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6767
    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
  6768
  note ab_fn = this
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6769
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6770
  have [simp]:"a = b" proof(rule ccontr)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6771
    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
  6772
    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
  6773
    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
  6774
      using lima limb unfolding LIMSEQ_def
93943da0a010 remove redundant lemma Lim_sequentially in favor of lemma LIMSEQ_def
huffman
parents: 44905
diff changeset
  6775
      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
  6776
    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
  6777
    have "dist (f (Suc n) x) (g a) \<le> dist (f n x) a"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6778
      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
  6779
    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
  6780
      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
  6781
    ultimately have "dist (f (Suc n) x) (g a) + dist (f (Suc n) y) (g b) < e" using n by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6782
    thus False unfolding e_def using ab_fn[of "Suc n"] by norm
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6783
  qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6784
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6785
  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
  6786
  { fix x y assume "x\<in>s" "y\<in>s" moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6787
    fix e::real assume "e>0" ultimately
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 44668
diff changeset
  6788
    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
  6789
  hence "continuous_on s g" unfolding continuous_on_iff by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6790
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6791
  hence "((snd \<circ> h \<circ> r) ---> g a) sequentially" unfolding continuous_on_sequentially
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6792
    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
  6793
    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
  6794
  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
  6795
    unfolding `a=b` and o_assoc by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6796
  moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6797
  { fix x assume "x\<in>s" "g x = x" "x\<noteq>a"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6798
    hence "False" using dist[THEN bspec[where x=a], THEN bspec[where x=x]]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6799
      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
  6800
  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
  6801
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6802
44131
5fc334b94e00 declare tendsto_const [intro] (accidentally removed in 230a8665c919)
huffman
parents: 44129
diff changeset
  6803
declare tendsto_const [intro] (* FIXME: move *)
5fc334b94e00 declare tendsto_const [intro] (accidentally removed in 230a8665c919)
huffman
parents: 44129
diff changeset
  6804
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6805
end