src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy
author hoelzl
Mon, 14 Jan 2013 17:29:04 +0100
changeset 50881 ae630bab13da
parent 50526 899c9c4e4a4c
child 50882 a382bf90867e
permissions -rw-r--r--
renamed countable_basis_space to second_countable_topology
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
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    11
  SEQ
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
  Linear_Algebra
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    15
  "~~/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
    16
  "~~/src/HOL/Library/FuncSet"
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
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
    24
lemma countable_rat: "countable \<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
    25
  unfolding Rats_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
    26
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    27
subsection {* Topological Basis *}
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    28
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    29
context topological_space
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    30
begin
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    31
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    32
definition "topological_basis B =
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    33
  ((\<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
    34
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    35
lemma topological_basis_iff:
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    36
  assumes "\<And>B'. B' \<in> B \<Longrightarrow> open B'"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    37
  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
    38
    (is "_ \<longleftrightarrow> ?rhs")
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    39
proof safe
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    40
  fix O' and x::'a
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    41
  assume H: "topological_basis B" "open O'" "x \<in> O'"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    42
  hence "(\<exists>B'\<subseteq>B. \<Union>B' = O')" by (simp add: topological_basis_def)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    43
  then obtain B' where "B' \<subseteq> B" "O' = \<Union>B'" by auto
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    44
  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
    45
next
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    46
  assume H: ?rhs
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    47
  show "topological_basis B" using assms unfolding topological_basis_def
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    48
  proof safe
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    49
    fix O'::"'a set" assume "open O'"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    50
    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
    51
      by (force intro: bchoice simp: Bex_def)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    52
    thus "\<exists>B'\<subseteq>B. \<Union>B' = O'"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    53
      by (auto intro: exI[where x="{f x |x. x \<in> O'}"])
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    54
  qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    55
qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    56
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    57
lemma topological_basisI:
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    58
  assumes "\<And>B'. B' \<in> B \<Longrightarrow> open B'"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    59
  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
    60
  shows "topological_basis B"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    61
  using assms by (subst topological_basis_iff) auto
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    62
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    63
lemma topological_basisE:
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    64
  fixes O'
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    65
  assumes "topological_basis B"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    66
  assumes "open O'"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    67
  assumes "x \<in> O'"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    68
  obtains B' where "B' \<in> B" "x \<in> B'" "B' \<subseteq> O'"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    69
proof atomize_elim
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    70
  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
    71
  with topological_basis_iff assms
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    72
  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
    73
qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    74
50094
84ddcf5364b4 allow arbitrary enumerations of basis in locale for generation of borel sets
immler
parents: 50087
diff changeset
    75
lemma topological_basis_open:
84ddcf5364b4 allow arbitrary enumerations of basis in locale for generation of borel sets
immler
parents: 50087
diff changeset
    76
  assumes "topological_basis B"
84ddcf5364b4 allow arbitrary enumerations of basis in locale for generation of borel sets
immler
parents: 50087
diff changeset
    77
  assumes "X \<in> B"
84ddcf5364b4 allow arbitrary enumerations of basis in locale for generation of borel sets
immler
parents: 50087
diff changeset
    78
  shows "open X"
84ddcf5364b4 allow arbitrary enumerations of basis in locale for generation of borel sets
immler
parents: 50087
diff changeset
    79
  using assms
84ddcf5364b4 allow arbitrary enumerations of basis in locale for generation of borel sets
immler
parents: 50087
diff changeset
    80
  by (simp add: topological_basis_def)
84ddcf5364b4 allow arbitrary enumerations of basis in locale for generation of borel sets
immler
parents: 50087
diff changeset
    81
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
    82
lemma basis_dense:
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
    83
  fixes B::"'a set set" and f::"'a set \<Rightarrow> 'a"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
    84
  assumes "topological_basis B"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
    85
  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
    86
  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
    87
proof (intro allI impI)
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
    88
  fix X::"'a set" assume "open X" "X \<noteq> {}"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
    89
  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
    90
  guess B' . note B' = this
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
    91
  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
    92
qed
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
    93
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    94
end
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
    95
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
    96
subsection {* Countable Basis *}
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
    97
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
    98
locale countable_basis =
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
    99
  fixes B::"'a::topological_space set set"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   100
  assumes is_basis: "topological_basis B"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   101
  assumes countable_basis: "countable B"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   102
begin
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   103
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   104
lemma open_countable_basis_ex:
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   105
  assumes "open X"
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   106
  shows "\<exists>B' \<subseteq> B. X = Union B'"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   107
  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
   108
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   109
lemma open_countable_basisE:
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   110
  assumes "open X"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   111
  obtains B' where "B' \<subseteq> B" "X = Union B'"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   112
  using assms open_countable_basis_ex by (atomize_elim) simp
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
lemma countable_dense_exists:
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   115
  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
   116
proof -
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   117
  let ?f = "(\<lambda>B'. SOME x. x \<in> B')"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   118
  have "countable (?f ` B)" using countable_basis by simp
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   119
  with basis_dense[OF is_basis, of ?f] show ?thesis
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   120
    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
   121
qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   122
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   123
lemma countable_dense_setE:
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   124
  obtains D :: "'a set"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   125
  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
   126
  using countable_dense_exists by blast
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   127
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   128
text {* Construction of an increasing sequence approximating open sets,
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   129
  therefore basis which is closed under union. *}
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   130
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   131
definition union_closed_basis::"'a set set" where
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   132
  "union_closed_basis = (\<lambda>l. \<Union>set l) ` lists B"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   133
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   134
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
   135
proof (rule topological_basisI)
84ddcf5364b4 allow arbitrary enumerations of basis in locale for generation of borel sets
immler
parents: 50087
diff changeset
   136
  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
   137
  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
   138
  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
   139
    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
   140
next
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   141
  fix B' assume "B' \<in> union_closed_basis"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   142
  thus "open B'"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   143
    using topological_basis_open[OF is_basis]
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   144
    by (auto simp: union_closed_basis_def)
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   145
qed
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
lemma countable_union_closed_basis: "countable union_closed_basis"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   148
  unfolding union_closed_basis_def using countable_basis by simp
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
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
   151
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   152
lemma union_closed_basis_ex:
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   153
 assumes X: "X \<in> union_closed_basis"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   154
 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
   155
proof -
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   156
  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
   157
  thus ?thesis by auto
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   158
qed
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   159
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   160
lemma union_closed_basisE:
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   161
  assumes "X \<in> union_closed_basis"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   162
  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
   163
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   164
lemma union_closed_basisI:
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   165
  assumes "finite B'" "X = \<Union>B'" "B' \<subseteq> B"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   166
  shows "X \<in> union_closed_basis"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   167
proof -
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   168
  from finite_list[OF `finite B'`] guess l ..
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   169
  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
   170
qed
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   171
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   172
lemma empty_basisI[intro]: "{} \<in> union_closed_basis"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   173
  by (rule union_closed_basisI[of "{}"]) auto
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   174
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   175
lemma union_basisI[intro]:
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   176
  assumes "X \<in> union_closed_basis" "Y \<in> union_closed_basis"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   177
  shows "X \<union> Y \<in> union_closed_basis"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   178
  using assms by (auto intro: union_closed_basisI elim!:union_closed_basisE)
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   179
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   180
lemma open_imp_Union_of_incseq:
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   181
  assumes "open X"
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   182
  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
   183
proof -
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   184
  from open_countable_basis_ex[OF `open X`]
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   185
  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
   186
  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
   187
  show ?thesis
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   188
  proof cases
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   189
    assume "B' \<noteq> {}"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   190
    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
   191
    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
   192
    have "incseq S" by (force simp: S_def incseq_Suc_iff)
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   193
    moreover
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   194
    have "(\<Union>j. S j) = X" unfolding B'
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   195
    proof safe
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   196
      fix x X assume "X \<in> B'" "x \<in> X"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   197
      then obtain n where "X = from_nat_into B' n"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   198
        by (metis `countable B'` from_nat_into_surj)
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   199
      also have "\<dots> \<subseteq> S n" by (auto simp: S_def)
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   200
      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
   201
    next
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   202
      fix x n
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   203
      assume "x \<in> S n"
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   204
      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
   205
        by (simp add: S_def)
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   206
      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
   207
      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
   208
      finally show "x \<in> \<Union>B'" .
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   209
    qed
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   210
    moreover have "range S \<subseteq> union_closed_basis" using B'
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   211
      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
   212
    ultimately show ?thesis by auto
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   213
  qed (auto simp: B')
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   214
qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   215
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   216
lemma open_incseqE:
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   217
  assumes "open X"
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   218
  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
   219
  using open_imp_Union_of_incseq assms by atomize_elim
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   220
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   221
end
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   222
50881
ae630bab13da renamed countable_basis_space to second_countable_topology
hoelzl
parents: 50526
diff changeset
   223
class second_countable_topology = topological_space +
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   224
  assumes ex_countable_basis:
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
   225
    "\<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
   226
50881
ae630bab13da renamed countable_basis_space to second_countable_topology
hoelzl
parents: 50526
diff changeset
   227
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
   228
  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
   229
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   230
subsection {* Polish spaces *}
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   231
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   232
text {* Textbooks define Polish spaces as completely metrizable.
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   233
  We assume the topology to be complete for a given metric. *}
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   234
50881
ae630bab13da renamed countable_basis_space to second_countable_topology
hoelzl
parents: 50526
diff changeset
   235
class polish_space = complete_space + second_countable_topology
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   236
44517
68e8eb0ce8aa minimize imports
huffman
parents: 44516
diff changeset
   237
subsection {* General notion of a topology as a value *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   238
44170
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   239
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
   240
typedef 'a topology = "{L::('a set) \<Rightarrow> bool. istopology L}"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   241
  morphisms "openin" "topology"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   242
  unfolding istopology_def by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   243
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   244
lemma istopology_open_in[intro]: "istopology(openin U)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   245
  using openin[of U] by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   246
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   247
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
   248
  using topology_inverse[unfolded mem_Collect_eq] .
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   249
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   250
lemma topology_inverse_iff: "istopology U \<longleftrightarrow> openin (topology U) = U"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   251
  using topology_inverse[of U] istopology_open_in[of "topology U"] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   252
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   253
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
   254
proof-
49711
e5aaae7eadc9 tuned proofs;
wenzelm
parents: 48125
diff changeset
   255
  { assume "T1=T2"
e5aaae7eadc9 tuned proofs;
wenzelm
parents: 48125
diff changeset
   256
    hence "\<forall>S. openin T1 S \<longleftrightarrow> openin T2 S" by simp }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   257
  moreover
49711
e5aaae7eadc9 tuned proofs;
wenzelm
parents: 48125
diff changeset
   258
  { 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
   259
    hence "openin T1 = openin T2" by (simp add: fun_eq_iff)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   260
    hence "topology (openin T1) = topology (openin T2)" by simp
49711
e5aaae7eadc9 tuned proofs;
wenzelm
parents: 48125
diff changeset
   261
    hence "T1 = T2" unfolding openin_inverse .
e5aaae7eadc9 tuned proofs;
wenzelm
parents: 48125
diff changeset
   262
  }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   263
  ultimately show ?thesis by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   264
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   265
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   266
text{* Infer the "universe" from union of all sets in the topology. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   267
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   268
definition "topspace T =  \<Union>{S. openin T S}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   269
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   270
subsubsection {* Main properties of open sets *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   271
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   272
lemma openin_clauses:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   273
  fixes U :: "'a topology"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   274
  shows "openin U {}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   275
  "\<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
   276
  "\<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
   277
  using openin[of U] unfolding istopology_def mem_Collect_eq
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   278
  by fast+
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   279
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   280
lemma openin_subset[intro]: "openin U S \<Longrightarrow> S \<subseteq> topspace U"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   281
  unfolding topspace_def by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   282
lemma openin_empty[simp]: "openin U {}" by (simp add: openin_clauses)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   283
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   284
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
   285
  using openin_clauses by simp
06475a1547cb fix lots of looping simp calls and other warnings
huffman
parents: 36360
diff changeset
   286
06475a1547cb fix lots of looping simp calls and other warnings
huffman
parents: 36360
diff changeset
   287
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
   288
  using openin_clauses by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   289
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   290
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
   291
  using openin_Union[of "{S,T}" U] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   292
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   293
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
   294
49711
e5aaae7eadc9 tuned proofs;
wenzelm
parents: 48125
diff changeset
   295
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
   296
  (is "?lhs \<longleftrightarrow> ?rhs")
36584
1535841fc2e9 prove lemma openin_subopen without using choice
huffman
parents: 36442
diff changeset
   297
proof
49711
e5aaae7eadc9 tuned proofs;
wenzelm
parents: 48125
diff changeset
   298
  assume ?lhs
e5aaae7eadc9 tuned proofs;
wenzelm
parents: 48125
diff changeset
   299
  then show ?rhs by auto
36584
1535841fc2e9 prove lemma openin_subopen without using choice
huffman
parents: 36442
diff changeset
   300
next
1535841fc2e9 prove lemma openin_subopen without using choice
huffman
parents: 36442
diff changeset
   301
  assume H: ?rhs
1535841fc2e9 prove lemma openin_subopen without using choice
huffman
parents: 36442
diff changeset
   302
  let ?t = "\<Union>{T. openin U T \<and> T \<subseteq> S}"
1535841fc2e9 prove lemma openin_subopen without using choice
huffman
parents: 36442
diff changeset
   303
  have "openin U ?t" by (simp add: openin_Union)
1535841fc2e9 prove lemma openin_subopen without using choice
huffman
parents: 36442
diff changeset
   304
  also have "?t = S" using H by auto
1535841fc2e9 prove lemma openin_subopen without using choice
huffman
parents: 36442
diff changeset
   305
  finally show "openin U S" .
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   306
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   307
49711
e5aaae7eadc9 tuned proofs;
wenzelm
parents: 48125
diff changeset
   308
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   309
subsubsection {* Closed sets *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   310
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   311
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
   312
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   313
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
   314
lemma closedin_empty[simp]: "closedin U {}" by (simp add: closedin_def)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   315
lemma closedin_topspace[intro,simp]:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   316
  "closedin U (topspace U)" by (simp add: closedin_def)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   317
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
   318
  by (auto simp add: Diff_Un closedin_def)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   319
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   320
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
   321
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
   322
  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
   323
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   324
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
   325
  using closedin_Inter[of "{S,T}" U] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   326
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   327
lemma Diff_Diff_Int: "A - (A - B) = A \<inter> B" by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   328
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
   329
  apply (auto simp add: closedin_def Diff_Diff_Int inf_absorb2)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   330
  apply (metis openin_subset subset_eq)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   331
  done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   332
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   333
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
   334
  by (simp add: openin_closedin_eq)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   335
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   336
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
   337
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   338
  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
   339
    by (auto simp add: topspace_def openin_subset)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   340
  then show ?thesis using oS cT by (auto simp add: closedin_def)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   341
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   342
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   343
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
   344
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   345
  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
   346
    by (auto simp add: topspace_def )
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   347
  then show ?thesis using oS cT by (auto simp add: openin_closedin_eq)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   348
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   349
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   350
subsubsection {* Subspace topology *}
44170
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   351
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   352
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
   353
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   354
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
   355
  (is "istopology ?L")
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   356
proof-
44170
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   357
  have "?L {}" by blast
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   358
  {fix A B assume A: "?L A" and B: "?L B"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   359
    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
   360
    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
   361
    then have "?L (A \<inter> B)" by blast}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   362
  moreover
44170
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   363
  {fix K assume K: "K \<subseteq> Collect ?L"
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   364
    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
   365
      apply (rule set_eqI)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   366
      apply (simp add: Ball_def image_iff)
44170
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   367
      by metis
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   368
    from K[unfolded th0 subset_image_iff]
44170
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   369
    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
   370
    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
   371
    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
   372
    ultimately have "?L (\<Union>K)" by blast}
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   373
  ultimately show ?thesis
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   374
    unfolding subset_eq mem_Collect_eq istopology_def by blast
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   375
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   376
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   377
lemma openin_subtopology:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   378
  "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
   379
  unfolding subtopology_def topology_inverse'[OF istopology_subtopology]
44170
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   380
  by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   381
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   382
lemma topspace_subtopology: "topspace(subtopology U V) = topspace U \<inter> V"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   383
  by (auto simp add: topspace_def openin_subtopology)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   384
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   385
lemma closedin_subtopology:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   386
  "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
   387
  unfolding closedin_def topspace_subtopology
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   388
  apply (simp add: openin_subtopology)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   389
  apply (rule iffI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   390
  apply clarify
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   391
  apply (rule_tac x="topspace U - T" in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   392
  by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   393
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   394
lemma openin_subtopology_refl: "openin (subtopology U V) V \<longleftrightarrow> V \<subseteq> topspace U"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   395
  unfolding openin_subtopology
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   396
  apply (rule iffI, clarify)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   397
  apply (frule openin_subset[of U])  apply blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   398
  apply (rule exI[where x="topspace U"])
49711
e5aaae7eadc9 tuned proofs;
wenzelm
parents: 48125
diff changeset
   399
  apply auto
e5aaae7eadc9 tuned proofs;
wenzelm
parents: 48125
diff changeset
   400
  done
e5aaae7eadc9 tuned proofs;
wenzelm
parents: 48125
diff changeset
   401
e5aaae7eadc9 tuned proofs;
wenzelm
parents: 48125
diff changeset
   402
lemma subtopology_superset:
e5aaae7eadc9 tuned proofs;
wenzelm
parents: 48125
diff changeset
   403
  assumes UV: "topspace U \<subseteq> V"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   404
  shows "subtopology U V = U"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   405
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   406
  {fix S
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   407
    {fix T assume T: "openin U T" "S = T \<inter> V"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   408
      from T openin_subset[OF T(1)] UV have eq: "S = T" by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   409
      have "openin U S" unfolding eq using T by blast}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   410
    moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   411
    {assume S: "openin U S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   412
      hence "\<exists>T. openin U T \<and> S = T \<inter> V"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   413
        using openin_subset[OF S] UV by auto}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   414
    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
   415
  then show ?thesis unfolding topology_eq openin_subtopology by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   416
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   417
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   418
lemma subtopology_topspace[simp]: "subtopology U (topspace U) = U"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   419
  by (simp add: subtopology_superset)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   420
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   421
lemma subtopology_UNIV[simp]: "subtopology U UNIV = U"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   422
  by (simp add: subtopology_superset)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   423
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   424
subsubsection {* The standard Euclidean topology *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   425
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   426
definition
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   427
  euclidean :: "'a::topological_space topology" where
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   428
  "euclidean = topology open"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   429
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   430
lemma open_openin: "open S \<longleftrightarrow> openin euclidean S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   431
  unfolding euclidean_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   432
  apply (rule cong[where x=S and y=S])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   433
  apply (rule topology_inverse[symmetric])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   434
  apply (auto simp add: istopology_def)
44170
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   435
  done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   436
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   437
lemma topspace_euclidean: "topspace euclidean = UNIV"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   438
  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
   439
  apply (rule set_eqI)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   440
  by (auto simp add: open_openin[symmetric])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   441
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   442
lemma topspace_euclidean_subtopology[simp]: "topspace (subtopology euclidean S) = S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   443
  by (simp add: topspace_euclidean topspace_subtopology)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   444
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   445
lemma closed_closedin: "closed S \<longleftrightarrow> closedin euclidean S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   446
  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
   447
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   448
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
   449
  by (simp add: open_openin openin_subopen[symmetric])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   450
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   451
text {* Basic "localization" results are handy for connectedness. *}
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   452
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   453
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
   454
  by (auto simp add: openin_subtopology open_openin[symmetric])
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   455
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   456
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
   457
  by (auto simp add: openin_open)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   458
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   459
lemma open_openin_trans[trans]:
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   460
 "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
   461
  by (metis Int_absorb1  openin_open_Int)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   462
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   463
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
   464
  by (auto simp add: openin_open)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   465
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   466
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
   467
  by (simp add: closedin_subtopology closed_closedin Int_ac)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   468
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   469
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
   470
  by (metis closedin_closed)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   471
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   472
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
   473
  apply (subgoal_tac "S \<inter> T = T" )
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   474
  apply auto
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   475
  apply (frule closedin_closed_Int[of T S])
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   476
  by simp
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   477
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   478
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
   479
  by (auto simp add: closedin_closed)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   480
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   481
lemma openin_euclidean_subtopology_iff:
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   482
  fixes S U :: "'a::metric_space set"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   483
  shows "openin (subtopology euclidean U) S
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   484
  \<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
   485
proof
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   486
  assume ?lhs thus ?rhs unfolding openin_open open_dist by blast
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   487
next
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   488
  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
   489
  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
   490
    unfolding T_def
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   491
    apply clarsimp
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   492
    apply (rule_tac x="d - dist x a" in exI)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   493
    apply (clarsimp simp add: less_diff_eq)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   494
    apply (erule rev_bexI)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   495
    apply (rule_tac x=d in exI, clarify)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   496
    apply (erule le_less_trans [OF dist_triangle])
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   497
    done
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   498
  assume ?rhs hence 2: "S = U \<inter> T"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   499
    unfolding T_def
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   500
    apply auto
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   501
    apply (drule (1) bspec, erule rev_bexI)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   502
    apply auto
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   503
    done
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   504
  from 1 2 show ?lhs
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   505
    unfolding openin_open open_dist by fast
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   506
qed
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   507
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   508
text {* These "transitivity" results are handy too *}
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   509
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   510
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
   511
  \<Longrightarrow> openin (subtopology euclidean U) S"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   512
  unfolding open_openin openin_open by blast
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   513
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   514
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
   515
  by (auto simp add: openin_open intro: openin_trans)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   516
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   517
lemma closedin_trans[trans]:
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   518
 "closedin (subtopology euclidean T) S \<Longrightarrow>
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   519
           closedin (subtopology euclidean U) T
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   520
           ==> closedin (subtopology euclidean U) S"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   521
  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
   522
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   523
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
   524
  by (auto simp add: closedin_closed intro: closedin_trans)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   525
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   526
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   527
subsection {* Open and closed balls *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   528
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   529
definition
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   530
  ball :: "'a::metric_space \<Rightarrow> real \<Rightarrow> 'a set" where
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   531
  "ball x e = {y. dist x y < e}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   532
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   533
definition
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   534
  cball :: "'a::metric_space \<Rightarrow> real \<Rightarrow> 'a set" where
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   535
  "cball x e = {y. dist x y \<le> e}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   536
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
   537
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
   538
  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
   539
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
   540
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
   541
  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
   542
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
   543
lemma mem_ball_0:
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   544
  fixes x :: "'a::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   545
  shows "x \<in> ball 0 e \<longleftrightarrow> norm x < e"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   546
  by (simp add: dist_norm)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   547
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
   548
lemma mem_cball_0:
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   549
  fixes x :: "'a::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   550
  shows "x \<in> cball 0 e \<longleftrightarrow> norm x \<le> e"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   551
  by (simp add: dist_norm)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   552
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
   553
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
   554
  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
   555
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
   556
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
   557
  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
   558
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   559
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
   560
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
   561
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
   562
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
   563
  by (simp add: set_eq_iff) arith
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   564
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   565
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
   566
  by (simp add: set_eq_iff)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   567
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   568
lemma diff_less_iff: "(a::real) - b > 0 \<longleftrightarrow> a > b"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   569
  "(a::real) - b < 0 \<longleftrightarrow> a < b"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   570
  "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
   571
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
   572
  "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
   573
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   574
lemma open_ball[intro, simp]: "open (ball x e)"
44170
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   575
  unfolding open_dist ball_def mem_Collect_eq Ball_def
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   576
  unfolding dist_commute
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   577
  apply clarify
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   578
  apply (rule_tac x="e - dist xa x" in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   579
  using dist_triangle_alt[where z=x]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   580
  apply (clarsimp simp add: diff_less_iff)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   581
  apply atomize
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   582
  apply (erule_tac x="y" in allE)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   583
  apply (erule_tac x="xa" in allE)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   584
  by arith
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   585
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   586
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
   587
  unfolding open_dist subset_eq mem_ball Ball_def dist_commute ..
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   588
33714
eb2574ac4173 Added new lemmas to Euclidean Space by Robert Himmelmann
hoelzl
parents: 33324
diff changeset
   589
lemma openE[elim?]:
eb2574ac4173 Added new lemmas to Euclidean Space by Robert Himmelmann
hoelzl
parents: 33324
diff changeset
   590
  assumes "open S" "x\<in>S" 
eb2574ac4173 Added new lemmas to Euclidean Space by Robert Himmelmann
hoelzl
parents: 33324
diff changeset
   591
  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
   592
  using assms unfolding open_contains_ball by auto
eb2574ac4173 Added new lemmas to Euclidean Space by Robert Himmelmann
hoelzl
parents: 33324
diff changeset
   593
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   594
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
   595
  by (metis open_contains_ball subset_eq centre_in_ball)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   596
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   597
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
   598
  unfolding mem_ball set_eq_iff
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   599
  apply (simp add: not_less)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   600
  by (metis zero_le_dist order_trans dist_self)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   601
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   602
lemma ball_empty[intro]: "e \<le> 0 ==> ball x e = {}" by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   603
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
   604
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
   605
  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
   606
  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
   607
  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
   608
  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
   609
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
   610
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
   611
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   612
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
   613
  fixes x :: "'a\<Colon>euclidean_space"
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   614
  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
   615
  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
   616
proof -
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   617
  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
   618
  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
   619
  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
   620
  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
   621
    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
   622
  qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   623
  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
   624
  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
   625
  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
   626
    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
   627
  qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   628
  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
   629
  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
   630
  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
   631
  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
   632
    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
   633
    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
   634
      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
   635
    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
   636
    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
   637
      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
   638
      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
   639
      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
   640
      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
   641
      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
   642
      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
   643
        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
   644
      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
   645
        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
   646
      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
   647
        by (simp add: power_divide)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   648
    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
   649
    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
   650
    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
   651
  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
   652
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
   653
 
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
   654
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
   655
  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
   656
  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
   657
  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
   658
  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
   659
  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
   660
  shows "M = (\<Union>f\<in>I. box (a' f) (b' f))"
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   661
proof safe
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   662
  fix x assume "x \<in> M"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
   663
  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
   664
    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
   665
  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
   666
    "\<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
   667
    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
   668
  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
   669
     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
   670
        (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
   671
qed (auto simp: I_def)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   672
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   673
subsection{* Connectedness *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   674
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   675
definition "connected S \<longleftrightarrow>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   676
  ~(\<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
   677
  \<and> ~(e1 \<inter> S = {}) \<and> ~(e2 \<inter> S = {}))"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   678
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   679
lemma connected_local:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   680
 "connected S \<longleftrightarrow> ~(\<exists>e1 e2.
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   681
                 openin (subtopology euclidean S) e1 \<and>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   682
                 openin (subtopology euclidean S) e2 \<and>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   683
                 S \<subseteq> e1 \<union> e2 \<and>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   684
                 e1 \<inter> e2 = {} \<and>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   685
                 ~(e1 = {}) \<and>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   686
                 ~(e2 = {}))"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   687
unfolding connected_def openin_open by (safe, blast+)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   688
34105
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
   689
lemma exists_diff:
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
   690
  fixes P :: "'a set \<Rightarrow> bool"
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
   691
  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
   692
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   693
  {assume "?lhs" hence ?rhs by blast }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   694
  moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   695
  {fix S assume H: "P S"
34105
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
   696
    have "S = - (- S)" by auto
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
   697
    with H have "P (- (- S))" by metis }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   698
  ultimately show ?thesis by metis
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   699
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   700
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   701
lemma connected_clopen: "connected S \<longleftrightarrow>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   702
        (\<forall>T. openin (subtopology euclidean S) T \<and>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   703
            closedin (subtopology euclidean S) T \<longrightarrow> T = {} \<or> T = S)" (is "?lhs \<longleftrightarrow> ?rhs")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   704
proof-
34105
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
   705
  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
   706
    unfolding connected_def openin_open closedin_closed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   707
    apply (subst exists_diff) by blast
34105
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
   708
  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
   709
    (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
   710
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   711
  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
   712
    (is "_ \<longleftrightarrow> \<not> (\<exists>t' t. ?Q t' t)")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   713
    unfolding connected_def openin_open closedin_closed by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   714
  {fix e2
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   715
    {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
   716
        by auto}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   717
    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
   718
  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
   719
  then show ?thesis unfolding th0 th1 by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   720
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   721
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   722
lemma connected_empty[simp, intro]: "connected {}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   723
  by (simp add: connected_def)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   724
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   725
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   726
subsection{* Limit points *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   727
44207
ea99698c2070 locale-ize some definitions, so perfect_space and heine_borel can inherit from the proper superclasses
huffman
parents: 44170
diff changeset
   728
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
   729
  islimpt:: "'a \<Rightarrow> 'a set \<Rightarrow> bool" (infixr "islimpt" 60) where
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   730
  "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
   731
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   732
lemma islimptI:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   733
  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
   734
  shows "x islimpt S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   735
  using assms unfolding islimpt_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   736
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   737
lemma islimptE:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   738
  assumes "x islimpt S" and "x \<in> T" and "open T"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   739
  obtains y where "y \<in> S" and "y \<in> T" and "y \<noteq> x"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   740
  using assms unfolding islimpt_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   741
44584
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
   742
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
   743
  unfolding islimpt_def eventually_at_topological by auto
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
   744
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
   745
lemma islimpt_subset: "\<lbrakk>x islimpt S; S \<subseteq> T\<rbrakk> \<Longrightarrow> x islimpt T"
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
   746
  unfolding islimpt_def by fast
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   747
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   748
lemma islimpt_approachable:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   749
  fixes x :: "'a::metric_space"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   750
  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
   751
  unfolding islimpt_iff_eventually eventually_at by fast
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   752
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   753
lemma islimpt_approachable_le:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   754
  fixes x :: "'a::metric_space"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   755
  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
   756
  unfolding islimpt_approachable
44584
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
   757
  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
   758
    THEN arg_cong [where f=Not]]
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
   759
  by (simp add: Bex_def conj_commute conj_left_commute)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   760
44571
bd91b77c4cd6 move class perfect_space into RealVector.thy;
huffman
parents: 44568
diff changeset
   761
lemma islimpt_UNIV_iff: "x islimpt UNIV \<longleftrightarrow> \<not> open {x}"
bd91b77c4cd6 move class perfect_space into RealVector.thy;
huffman
parents: 44568
diff changeset
   762
  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
   763
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   764
text {* A perfect space has no isolated points. *}
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   765
44571
bd91b77c4cd6 move class perfect_space into RealVector.thy;
huffman
parents: 44568
diff changeset
   766
lemma islimpt_UNIV [simp, intro]: "(x::'a::perfect_space) islimpt UNIV"
bd91b77c4cd6 move class perfect_space into RealVector.thy;
huffman
parents: 44568
diff changeset
   767
  unfolding islimpt_UNIV_iff by (rule not_open_singleton)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   768
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   769
lemma perfect_choose_dist:
44072
5b970711fb39 class perfect_space inherits from topological_space;
huffman
parents: 43338
diff changeset
   770
  fixes x :: "'a::{perfect_space, metric_space}"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   771
  shows "0 < r \<Longrightarrow> \<exists>a. a \<noteq> x \<and> dist a x < r"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   772
using islimpt_UNIV [of x]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   773
by (simp add: islimpt_approachable)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   774
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   775
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
   776
  unfolding closed_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   777
  apply (subst open_subopen)
34105
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
   778
  apply (simp add: islimpt_def subset_eq)
44170
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44167
diff changeset
   779
  by (metis ComplE ComplI)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   780
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   781
lemma islimpt_EMPTY[simp]: "\<not> x islimpt {}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   782
  unfolding islimpt_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   783
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   784
lemma finite_set_avoid:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   785
  fixes a :: "'a::metric_space"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   786
  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
   787
proof(induct rule: finite_induct[OF fS])
41863
e5104b436ea1 removed dependency on Dense_Linear_Order
boehmes
parents: 41413
diff changeset
   788
  case 1 thus ?case by (auto intro: zero_less_one)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   789
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   790
  case (2 x F)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   791
  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
   792
  {assume "x = a" hence ?case using d by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   793
  moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   794
  {assume xa: "x\<noteq>a"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   795
    let ?d = "min d (dist a x)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   796
    have dp: "?d > 0" using xa d(1) using dist_nz by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   797
    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
   798
    with dp xa have ?case by(auto intro!: exI[where x="?d"]) }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   799
  ultimately show ?case by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   800
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   801
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   802
lemma islimpt_finite:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   803
  fixes S :: "'a::metric_space set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   804
  assumes fS: "finite S" shows "\<not> a islimpt S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   805
  unfolding islimpt_approachable
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   806
  using finite_set_avoid[OF fS, of a] by (metis dist_commute  not_le)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   807
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   808
lemma islimpt_Un: "x islimpt (S \<union> T) \<longleftrightarrow> x islimpt S \<or> x islimpt T"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   809
  apply (rule iffI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   810
  defer
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   811
  apply (metis Un_upper1 Un_upper2 islimpt_subset)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   812
  unfolding islimpt_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   813
  apply (rule ccontr, clarsimp, rename_tac A B)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   814
  apply (drule_tac x="A \<inter> B" in spec)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   815
  apply (auto simp add: open_Int)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   816
  done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   817
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   818
lemma discrete_imp_closed:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   819
  fixes S :: "'a::metric_space set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   820
  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
   821
  shows "closed S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   822
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   823
  {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
   824
    from e have e2: "e/2 > 0" by arith
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   825
    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
   826
    let ?m = "min (e/2) (dist x y) "
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   827
    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
   828
    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
   829
    have th: "dist z y < e" using z y
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   830
      by (intro dist_triangle_lt [where z=x], simp)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   831
    from d[rule_format, OF y(1) z(1) th] y z
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   832
    have False by (auto simp add: dist_commute)}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   833
  then show ?thesis by (metis islimpt_approachable closed_limpt [where 'a='a])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   834
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   835
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   836
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   837
subsection {* Interior of a Set *}
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   838
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   839
definition "interior S = \<Union>{T. open T \<and> T \<subseteq> S}"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   840
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   841
lemma interiorI [intro?]:
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   842
  assumes "open T" and "x \<in> T" and "T \<subseteq> S"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   843
  shows "x \<in> interior S"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   844
  using assms unfolding interior_def by fast
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   845
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   846
lemma interiorE [elim?]:
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   847
  assumes "x \<in> interior S"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   848
  obtains T where "open T" and "x \<in> T" and "T \<subseteq> S"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   849
  using assms unfolding interior_def by fast
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   850
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   851
lemma open_interior [simp, intro]: "open (interior S)"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   852
  by (simp add: interior_def open_Union)
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   853
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   854
lemma interior_subset: "interior S \<subseteq> S"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   855
  by (auto simp add: interior_def)
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   856
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   857
lemma interior_maximal: "T \<subseteq> S \<Longrightarrow> open T \<Longrightarrow> T \<subseteq> interior S"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   858
  by (auto simp add: interior_def)
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   859
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   860
lemma interior_open: "open S \<Longrightarrow> interior S = S"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   861
  by (intro equalityI interior_subset interior_maximal subset_refl)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   862
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   863
lemma interior_eq: "interior S = S \<longleftrightarrow> open S"
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   864
  by (metis open_interior interior_open)
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   865
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   866
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
   867
  by (metis interior_maximal interior_subset subset_trans)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   868
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   869
lemma interior_empty [simp]: "interior {} = {}"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   870
  using open_empty by (rule interior_open)
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   871
44522
2f7e9d890efe rename subset_{interior,closure} to {interior,closure}_mono;
huffman
parents: 44519
diff changeset
   872
lemma interior_UNIV [simp]: "interior UNIV = UNIV"
2f7e9d890efe rename subset_{interior,closure} to {interior,closure}_mono;
huffman
parents: 44519
diff changeset
   873
  using open_UNIV by (rule interior_open)
2f7e9d890efe rename subset_{interior,closure} to {interior,closure}_mono;
huffman
parents: 44519
diff changeset
   874
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   875
lemma interior_interior [simp]: "interior (interior S) = interior S"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   876
  using open_interior by (rule interior_open)
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   877
44522
2f7e9d890efe rename subset_{interior,closure} to {interior,closure}_mono;
huffman
parents: 44519
diff changeset
   878
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
   879
  by (auto simp add: interior_def)
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   880
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   881
lemma interior_unique:
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   882
  assumes "T \<subseteq> S" and "open T"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   883
  assumes "\<And>T'. T' \<subseteq> S \<Longrightarrow> open T' \<Longrightarrow> T' \<subseteq> T"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   884
  shows "interior S = T"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   885
  by (intro equalityI assms interior_subset open_interior interior_maximal)
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   886
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   887
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
   888
  by (intro equalityI Int_mono Int_greatest interior_mono Int_lower1
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   889
    Int_lower2 interior_maximal interior_subset open_Int open_interior)
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   890
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   891
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
   892
  using open_contains_ball_eq [where S="interior S"]
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   893
  by (simp add: open_subset_interior)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   894
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   895
lemma interior_limit_point [intro]:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   896
  fixes x :: "'a::perfect_space"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   897
  assumes x: "x \<in> interior S" shows "x islimpt S"
44072
5b970711fb39 class perfect_space inherits from topological_space;
huffman
parents: 43338
diff changeset
   898
  using x islimpt_UNIV [of x]
5b970711fb39 class perfect_space inherits from topological_space;
huffman
parents: 43338
diff changeset
   899
  unfolding interior_def islimpt_def
5b970711fb39 class perfect_space inherits from topological_space;
huffman
parents: 43338
diff changeset
   900
  apply (clarsimp, rename_tac T T')
5b970711fb39 class perfect_space inherits from topological_space;
huffman
parents: 43338
diff changeset
   901
  apply (drule_tac x="T \<inter> T'" in spec)
5b970711fb39 class perfect_space inherits from topological_space;
huffman
parents: 43338
diff changeset
   902
  apply (auto simp add: open_Int)
5b970711fb39 class perfect_space inherits from topological_space;
huffman
parents: 43338
diff changeset
   903
  done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   904
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   905
lemma interior_closed_Un_empty_interior:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   906
  assumes cS: "closed S" and iT: "interior T = {}"
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   907
  shows "interior (S \<union> T) = interior S"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   908
proof
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   909
  show "interior S \<subseteq> interior (S \<union> T)"
44522
2f7e9d890efe rename subset_{interior,closure} to {interior,closure}_mono;
huffman
parents: 44519
diff changeset
   910
    by (rule interior_mono, rule Un_upper1)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   911
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   912
  show "interior (S \<union> T) \<subseteq> interior S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   913
  proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   914
    fix x assume "x \<in> interior (S \<union> T)"
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   915
    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
   916
    show "x \<in> interior S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   917
    proof (rule ccontr)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   918
      assume "x \<notin> interior S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   919
      with `x \<in> R` `open R` obtain y where "y \<in> R - S"
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   920
        unfolding interior_def by fast
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   921
      from `open R` `closed S` have "open (R - S)" by (rule open_Diff)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   922
      from `R \<subseteq> S \<union> T` have "R - S \<subseteq> T" by fast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   923
      from `y \<in> R - S` `open (R - S)` `R - S \<subseteq> T` `interior T = {}`
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   924
      show "False" unfolding interior_def by fast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   925
    qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   926
  qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   927
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   928
44365
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
   929
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
   930
proof (rule interior_unique)
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
   931
  show "interior A \<times> interior B \<subseteq> A \<times> B"
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
   932
    by (intro Sigma_mono interior_subset)
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
   933
  show "open (interior A \<times> interior B)"
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
   934
    by (intro open_Times open_interior)
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   935
  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
   936
  proof (safe)
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   937
    fix x y assume "(x, y) \<in> T"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   938
    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
   939
      using `open T` unfolding open_prod_def by fast
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   940
    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
   941
      using `T \<subseteq> A \<times> B` by auto
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   942
    thus "x \<in> interior A" and "y \<in> interior B"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   943
      by (auto intro: interiorI)
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   944
  qed
44365
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
   945
qed
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
   946
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   947
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
   948
subsection {* Closure of a Set *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   949
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   950
definition "closure S = S \<union> {x | x. x islimpt S}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   951
44518
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
   952
lemma interior_closure: "interior S = - (closure (- S))"
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
   953
  unfolding interior_def closure_def islimpt_def by auto
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
   954
34105
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
   955
lemma closure_interior: "closure S = - interior (- S)"
44518
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
   956
  unfolding interior_closure by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   957
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   958
lemma closed_closure[simp, intro]: "closed (closure S)"
44518
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
   959
  unfolding closure_interior by (simp add: closed_Compl)
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
   960
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
   961
lemma closure_subset: "S \<subseteq> closure S"
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
   962
  unfolding closure_def by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   963
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   964
lemma closure_hull: "closure S = closed hull S"
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   965
  unfolding hull_def closure_interior interior_def by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   966
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   967
lemma closure_eq: "closure S = S \<longleftrightarrow> closed S"
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   968
  unfolding closure_hull using closed_Inter by (rule hull_eq)
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   969
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   970
lemma closure_closed [simp]: "closed S \<Longrightarrow> closure S = S"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   971
  unfolding closure_eq .
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   972
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   973
lemma closure_closure [simp]: "closure (closure S) = closure S"
44518
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
   974
  unfolding closure_hull by (rule hull_hull)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   975
44522
2f7e9d890efe rename subset_{interior,closure} to {interior,closure}_mono;
huffman
parents: 44519
diff changeset
   976
lemma closure_mono: "S \<subseteq> T \<Longrightarrow> closure S \<subseteq> closure T"
44518
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
   977
  unfolding closure_hull by (rule hull_mono)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   978
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   979
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
   980
  unfolding closure_hull by (rule hull_minimal)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   981
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   982
lemma closure_unique:
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   983
  assumes "S \<subseteq> T" and "closed T"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   984
  assumes "\<And>T'. S \<subseteq> T' \<Longrightarrow> closed T' \<Longrightarrow> T \<subseteq> T'"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   985
  shows "closure S = T"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   986
  using assms unfolding closure_hull by (rule hull_unique)
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   987
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
   988
lemma closure_empty [simp]: "closure {} = {}"
44518
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
   989
  using closed_empty by (rule closure_closed)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   990
44522
2f7e9d890efe rename subset_{interior,closure} to {interior,closure}_mono;
huffman
parents: 44519
diff changeset
   991
lemma closure_UNIV [simp]: "closure UNIV = UNIV"
44518
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
   992
  using closed_UNIV by (rule closure_closed)
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
   993
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
   994
lemma closure_union [simp]: "closure (S \<union> T) = closure S \<union> closure T"
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
   995
  unfolding closure_interior by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   996
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   997
lemma closure_eq_empty: "closure S = {} \<longleftrightarrow> S = {}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   998
  using closure_empty closure_subset[of S]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   999
  by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1000
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1001
lemma closure_subset_eq: "closure S \<subseteq> S \<longleftrightarrow> closed S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1002
  using closure_eq[of S] closure_subset[of S]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1003
  by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1004
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1005
lemma open_inter_closure_eq_empty:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1006
  "open S \<Longrightarrow> (S \<inter> closure T) = {} \<longleftrightarrow> S \<inter> T = {}"
34105
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
  1007
  using open_subset_interior[of S "- T"]
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
  1008
  using interior_subset[of "- T"]
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1009
  unfolding closure_interior
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1010
  by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1011
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1012
lemma open_inter_closure_subset:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1013
  "open S \<Longrightarrow> (S \<inter> (closure T)) \<subseteq> closure(S \<inter> T)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1014
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1015
  fix x
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1016
  assume as: "open S" "x \<in> S \<inter> closure T"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1017
  { assume *:"x islimpt T"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1018
    have "x islimpt (S \<inter> T)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1019
    proof (rule islimptI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1020
      fix A
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1021
      assume "x \<in> A" "open A"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1022
      with as have "x \<in> A \<inter> S" "open (A \<inter> S)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1023
        by (simp_all add: open_Int)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1024
      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
  1025
        by (rule islimptE)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1026
      hence "y \<in> S \<inter> T" "y \<in> A \<and> y \<noteq> x"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1027
        by simp_all
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1028
      thus "\<exists>y\<in>(S \<inter> T). y \<in> A \<and> y \<noteq> x" ..
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1029
    qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1030
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1031
  then show "x \<in> closure (S \<inter> T)" using as
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1032
    unfolding closure_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1033
    by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1034
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1035
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1036
lemma closure_complement: "closure (- S) = - interior S"
44518
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
  1037
  unfolding closure_interior by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1038
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1039
lemma interior_complement: "interior (- S) = - closure S"
44518
219a6fe4cfae add lemma closure_union;
huffman
parents: 44517
diff changeset
  1040
  unfolding closure_interior by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1041
44365
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1042
lemma closure_Times: "closure (A \<times> B) = closure A \<times> closure B"
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1043
proof (rule closure_unique)
44365
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1044
  show "A \<times> B \<subseteq> closure A \<times> closure B"
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1045
    by (intro Sigma_mono closure_subset)
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1046
  show "closed (closure A \<times> closure B)"
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1047
    by (intro closed_Times closed_closure)
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1048
  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
  1049
    apply (simp add: closed_def open_prod_def, clarify)
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1050
    apply (rule ccontr)
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1051
    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
  1052
    apply (simp add: closure_interior interior_def)
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1053
    apply (drule_tac x=C in spec)
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1054
    apply (drule_tac x=D in spec)
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1055
    apply auto
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1056
    done
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1057
qed
5daa55003649 add lemmas interior_Times and closure_Times
huffman
parents: 44342
diff changeset
  1058
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1059
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1060
subsection {* Frontier (aka boundary) *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1061
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1062
definition "frontier S = closure S - interior S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1063
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1064
lemma frontier_closed: "closed(frontier S)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1065
  by (simp add: frontier_def closed_Diff)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1066
34105
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
  1067
lemma frontier_closures: "frontier S = (closure S) \<inter> (closure(- S))"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1068
  by (auto simp add: frontier_def interior_closure)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1069
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1070
lemma frontier_straddle:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1071
  fixes a :: "'a::metric_space"
44909
1f5d6eb73549 shorten proof of frontier_straddle
huffman
parents: 44907
diff changeset
  1072
  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
  1073
  unfolding frontier_def closure_interior
1f5d6eb73549 shorten proof of frontier_straddle
huffman
parents: 44907
diff changeset
  1074
  by (auto simp add: mem_interior subset_eq ball_def)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1075
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1076
lemma frontier_subset_closed: "closed S \<Longrightarrow> frontier S \<subseteq> S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1077
  by (metis frontier_def closure_closed Diff_subset)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1078
34964
4e8be3c04d37 Replaced vec1 and dest_vec1 by abbreviation.
hoelzl
parents: 34291
diff changeset
  1079
lemma frontier_empty[simp]: "frontier {} = {}"
36362
06475a1547cb fix lots of looping simp calls and other warnings
huffman
parents: 36360
diff changeset
  1080
  by (simp add: frontier_def)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1081
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1082
lemma frontier_subset_eq: "frontier S \<subseteq> S \<longleftrightarrow> closed S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1083
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1084
  { assume "frontier S \<subseteq> S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1085
    hence "closure S \<subseteq> S" using interior_subset unfolding frontier_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1086
    hence "closed S" using closure_subset_eq by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1087
  }
36362
06475a1547cb fix lots of looping simp calls and other warnings
huffman
parents: 36360
diff changeset
  1088
  thus ?thesis using frontier_subset_closed[of S] ..
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1089
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1090
34105
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
  1091
lemma frontier_complement: "frontier(- S) = frontier S"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1092
  by (auto simp add: frontier_def closure_complement interior_complement)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1093
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1094
lemma frontier_disjoint_eq: "frontier S \<inter> S = {} \<longleftrightarrow> open S"
34105
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
  1095
  using frontier_complement frontier_subset_eq[of "- S"]
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
  1096
  unfolding open_closed by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1097
44081
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44076
diff changeset
  1098
subsection {* Filters and the ``eventually true'' quantifier *}
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44076
diff changeset
  1099
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1100
definition
44081
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44076
diff changeset
  1101
  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
  1102
    (infixr "indirection" 70) where
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1103
  "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
  1104
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  1105
text {* Identify Trivial limits, where we can't approach arbitrarily closely. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1106
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1107
lemma trivial_limit_within:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1108
  shows "trivial_limit (at a within S) \<longleftrightarrow> \<not> a islimpt S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1109
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1110
  assume "trivial_limit (at a within S)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1111
  thus "\<not> a islimpt S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1112
    unfolding trivial_limit_def
36358
246493d61204 define nets directly as filters, instead of as filter bases
huffman
parents: 36336
diff changeset
  1113
    unfolding eventually_within eventually_at_topological
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1114
    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
  1115
    apply (clarsimp simp add: set_eq_iff)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1116
    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
  1117
    apply (clarsimp, drule_tac x=y in bspec, simp_all)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1118
    done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1119
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1120
  assume "\<not> a islimpt S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1121
  thus "trivial_limit (at a within S)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1122
    unfolding trivial_limit_def
36358
246493d61204 define nets directly as filters, instead of as filter bases
huffman
parents: 36336
diff changeset
  1123
    unfolding eventually_within eventually_at_topological
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1124
    unfolding islimpt_def
36358
246493d61204 define nets directly as filters, instead of as filter bases
huffman
parents: 36336
diff changeset
  1125
    apply clarsimp
246493d61204 define nets directly as filters, instead of as filter bases
huffman
parents: 36336
diff changeset
  1126
    apply (rule_tac x=T in exI)
246493d61204 define nets directly as filters, instead of as filter bases
huffman
parents: 36336
diff changeset
  1127
    apply auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1128
    done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1129
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1130
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1131
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
  1132
  using trivial_limit_within [of a UNIV] by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1133
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1134
lemma trivial_limit_at:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1135
  fixes a :: "'a::perfect_space"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1136
  shows "\<not> trivial_limit (at a)"
44571
bd91b77c4cd6 move class perfect_space into RealVector.thy;
huffman
parents: 44568
diff changeset
  1137
  by (rule at_neq_bot)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1138
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1139
lemma trivial_limit_at_infinity:
44081
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44076
diff changeset
  1140
  "\<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
  1141
  unfolding trivial_limit_def eventually_at_infinity
246493d61204 define nets directly as filters, instead of as filter bases
huffman
parents: 36336
diff changeset
  1142
  apply clarsimp
44072
5b970711fb39 class perfect_space inherits from topological_space;
huffman
parents: 43338
diff changeset
  1143
  apply (subgoal_tac "\<exists>x::'a. x \<noteq> 0", clarify)
5b970711fb39 class perfect_space inherits from topological_space;
huffman
parents: 43338
diff changeset
  1144
   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
  1145
  apply (cut_tac islimpt_UNIV [of "0::'a", unfolded islimpt_def])
5b970711fb39 class perfect_space inherits from topological_space;
huffman
parents: 43338
diff changeset
  1146
  apply (drule_tac x=UNIV in spec, simp)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1147
  done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1148
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  1149
text {* Some property holds "sufficiently close" to the limit point. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1150
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1151
lemma eventually_at: (* FIXME: this replaces Limits.eventually_at *)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1152
  "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
  1153
unfolding eventually_at dist_nz by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1154
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  1155
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
  1156
  "eventually P (at a within S) \<longleftrightarrow>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1157
        (\<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
  1158
  by (rule eventually_within_less)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1159
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1160
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
  1161
  unfolding trivial_limit_def
246493d61204 define nets directly as filters, instead of as filter bases
huffman
parents: 36336
diff changeset
  1162
  by (auto elim: eventually_rev_mp)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1163
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1164
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
  1165
  by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1166
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1167
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
  1168
  by (simp add: filter_eq_iff)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1169
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1170
text{* Combining theorems for "eventually" *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1171
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1172
lemma eventually_rev_mono:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1173
  "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
  1174
using eventually_mono [of P Q] by fast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1176
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
  1177
  by (simp add: eventually_False)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1178
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1179
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  1180
subsection {* Limits *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1181
44081
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44076
diff changeset
  1182
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
  1183
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44076
diff changeset
  1184
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
  1185
  where "Lim A f = (THE l. (f ---> l) A)"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1186
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1187
lemma Lim:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1188
 "(f ---> l) net \<longleftrightarrow>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1189
        trivial_limit net \<or>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1190
        (\<forall>e>0. eventually (\<lambda>x. dist (f x) l < e) net)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1191
  unfolding tendsto_iff trivial_limit_eq by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1192
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1193
text{* Show that they yield usual definitions in the various cases. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1194
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1195
lemma Lim_within_le: "(f ---> l)(at a within S) \<longleftrightarrow>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1196
           (\<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
  1197
  by (auto simp add: tendsto_iff eventually_within_le)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1198
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1199
lemma Lim_within: "(f ---> l) (at a within S) \<longleftrightarrow>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1200
        (\<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
  1201
  by (auto simp add: tendsto_iff eventually_within)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1202
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1203
lemma Lim_at: "(f ---> l) (at a) \<longleftrightarrow>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1204
        (\<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
  1205
  by (auto simp add: tendsto_iff eventually_at)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1206
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1207
lemma Lim_at_infinity:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1208
  "(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
  1209
  by (auto simp add: tendsto_iff eventually_at_infinity)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1210
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1211
lemma Lim_eventually: "eventually (\<lambda>x. f x = l) net \<Longrightarrow> (f ---> l) net"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1212
  by (rule topological_tendstoI, auto elim: eventually_rev_mono)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1213
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1214
text{* The expected monotonicity property. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1215
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1216
lemma Lim_within_empty: "(f ---> l) (net within {})"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1217
  unfolding tendsto_def Limits.eventually_within by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1218
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1219
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
  1220
  unfolding tendsto_def Limits.eventually_within
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1221
  by (auto elim!: eventually_elim1)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1222
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1223
lemma Lim_Un: assumes "(f ---> l) (net within S)" "(f ---> l) (net within T)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1224
  shows "(f ---> l) (net within (S \<union> T))"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1225
  using assms unfolding tendsto_def Limits.eventually_within
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1226
  apply clarify
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1227
  apply (drule spec, drule (1) mp, drule (1) mp)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1228
  apply (drule spec, drule (1) mp, drule (1) mp)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1229
  apply (auto elim: eventually_elim2)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1230
  done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1231
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1232
lemma Lim_Un_univ:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1233
 "(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
  1234
        ==> (f ---> l) net"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1235
  by (metis Lim_Un within_UNIV)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1236
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1237
text{* Interrelations between restricted and unrestricted limits. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1238
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1239
lemma Lim_at_within: "(f ---> l) net ==> (f ---> l)(net within S)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1240
  (* FIXME: rename *)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1241
  unfolding tendsto_def Limits.eventually_within
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1242
  apply (clarify, drule spec, drule (1) mp, drule (1) mp)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1243
  by (auto elim!: eventually_elim1)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1244
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1245
lemma eventually_within_interior:
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1246
  assumes "x \<in> interior S"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1247
  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
  1248
proof-
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  1249
  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
  1250
  { assume "?lhs"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1251
    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
  1252
      unfolding Limits.eventually_within Limits.eventually_at_topological
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1253
      by auto
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1254
    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
  1255
      by auto
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1256
    then have "?rhs"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1257
      unfolding Limits.eventually_at_topological by auto
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1258
  } moreover
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1259
  { assume "?rhs" hence "?lhs"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1260
      unfolding Limits.eventually_within
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1261
      by (auto elim: eventually_elim1)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1262
  } ultimately
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1263
  show "?thesis" ..
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1264
qed
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1265
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1266
lemma at_within_interior:
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1267
  "x \<in> interior S \<Longrightarrow> at x within S = at x"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1268
  by (simp add: filter_eq_iff eventually_within_interior)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1269
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1270
lemma at_within_open:
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1271
  "\<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
  1272
  by (simp only: at_within_interior interior_open)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1273
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1274
lemma Lim_within_open:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1275
  fixes f :: "'a::topological_space \<Rightarrow> 'b::topological_space"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1276
  assumes"a \<in> S" "open S"
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1277
  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
  1278
  using assms by (simp only: at_within_open)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1279
43338
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1280
lemma Lim_within_LIMSEQ:
44584
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  1281
  fixes a :: "'a::metric_space"
43338
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1282
  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
  1283
  shows "(X ---> L) (at a within T)"
44584
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  1284
  using assms unfolding tendsto_def [where l=L]
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  1285
  by (simp add: sequentially_imp_eventually_within)
43338
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1286
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1287
lemma Lim_right_bound:
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1288
  fixes f :: "real \<Rightarrow> real"
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1289
  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
  1290
  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
  1291
  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
  1292
proof cases
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1293
  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
  1294
next
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1295
  assume [simp]: "{x<..} \<inter> I \<noteq> {}"
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1296
  show ?thesis
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1297
  proof (rule Lim_within_LIMSEQ, safe)
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1298
    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
  1299
    
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1300
    show "(\<lambda>n. f (S n)) ----> Inf (f ` ({x<..} \<inter> I))"
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1301
    proof (rule LIMSEQ_I, rule ccontr)
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1302
      fix r :: real assume "0 < r"
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1303
      with Inf_close[of "f ` ({x<..} \<inter> I)" r]
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1304
      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
  1305
      from `x < y` have "0 < y - x" by auto
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1306
      from S(2)[THEN LIMSEQ_D, OF this]
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1307
      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
  1308
      
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1309
      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
  1310
      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
  1311
        using S bnd by (intro Inf_lower[where z=K]) auto
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1312
      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
  1313
        by (auto simp: not_less field_simps)
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1314
      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
  1315
      show False by auto
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1316
    qed
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1317
  qed
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1318
qed
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1319
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1320
text{* Another limit point characterization. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1321
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1322
lemma islimpt_sequential:
36667
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1323
  fixes x :: "'a::metric_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1324
  shows "x islimpt S \<longleftrightarrow> (\<exists>f. (\<forall>n::nat. f n \<in> S -{x}) \<and> (f ---> x) sequentially)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1325
    (is "?lhs = ?rhs")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1326
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1327
  assume ?lhs
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1328
  then obtain f where f:"\<forall>y. y>0 \<longrightarrow> f y \<in> S \<and> f y \<noteq> x \<and> dist (f y) x < y"
44584
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  1329
    unfolding islimpt_approachable
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  1330
    using choice[of "\<lambda>e y. e>0 \<longrightarrow> y\<in>S \<and> y\<noteq>x \<and> dist y x < e"] by auto
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  1331
  let ?I = "\<lambda>n. inverse (real (Suc n))"
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  1332
  have "\<forall>n. f (?I n) \<in> S - {x}" using f by simp
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  1333
  moreover have "(\<lambda>n. f (?I n)) ----> x"
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  1334
  proof (rule metric_tendsto_imp_tendsto)
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  1335
    show "?I ----> 0"
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  1336
      by (rule LIMSEQ_inverse_real_of_nat)
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  1337
    show "eventually (\<lambda>n. dist (f (?I n)) x \<le> dist (?I n) 0) sequentially"
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  1338
      by (simp add: norm_conv_dist [symmetric] less_imp_le f)
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  1339
  qed
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  1340
  ultimately show ?rhs by fast
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1341
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1342
  assume ?rhs
44907
93943da0a010 remove redundant lemma Lim_sequentially in favor of lemma LIMSEQ_def
huffman
parents: 44905
diff changeset
  1343
  then obtain f::"nat\<Rightarrow>'a"  where f:"(\<forall>n. f n \<in> S - {x})" "(\<forall>e>0. \<exists>N. \<forall>n\<ge>N. dist (f n) x < e)" unfolding LIMSEQ_def by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1344
  { fix e::real assume "e>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1345
    then obtain N where "dist (f N) x < e" using f(2) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1346
    moreover have "f N\<in>S" "f N \<noteq> x" using f(1) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1347
    ultimately have "\<exists>x'\<in>S. x' \<noteq> x \<and> dist x' x < e" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1348
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1349
  thus ?lhs unfolding islimpt_approachable by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1350
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1351
44125
230a8665c919 mark some redundant theorems as legacy
huffman
parents: 44122
diff changeset
  1352
lemma Lim_inv: (* TODO: delete *)
44081
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44076
diff changeset
  1353
  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
  1354
  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
  1355
  shows "((inverse o f) ---> inverse l) A"
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  1356
  unfolding o_def using assms by (rule tendsto_inverse)
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  1357
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1358
lemma Lim_null:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1359
  fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
44125
230a8665c919 mark some redundant theorems as legacy
huffman
parents: 44122
diff changeset
  1360
  shows "(f ---> l) net \<longleftrightarrow> ((\<lambda>x. f(x) - l) ---> 0) net"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1361
  by (simp add: Lim dist_norm)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1362
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1363
lemma Lim_null_comparison:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1364
  fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1365
  assumes "eventually (\<lambda>x. norm (f x) \<le> g x) net" "(g ---> 0) net"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1366
  shows "(f ---> 0) net"
44252
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1367
proof (rule metric_tendsto_imp_tendsto)
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1368
  show "(g ---> 0) net" by fact
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1369
  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
  1370
    using assms(1) by (rule eventually_elim1, simp add: dist_norm)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1371
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1372
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1373
lemma Lim_transform_bound:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1374
  fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1375
  fixes g :: "'a \<Rightarrow> 'c::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1376
  assumes "eventually (\<lambda>n. norm(f n) <= norm(g n)) net"  "(g ---> 0) net"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1377
  shows "(f ---> 0) net"
44252
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1378
  using assms(1) tendsto_norm_zero [OF assms(2)]
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1379
  by (rule Lim_null_comparison)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1380
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1381
text{* Deducing things about the limit from the elements. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1382
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1383
lemma Lim_in_closed_set:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1384
  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
  1385
  shows "l \<in> S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1386
proof (rule ccontr)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1387
  assume "l \<notin> S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1388
  with `closed S` have "open (- S)" "l \<in> - S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1389
    by (simp_all add: open_Compl)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1390
  with assms(4) have "eventually (\<lambda>x. f x \<in> - S) net"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1391
    by (rule topological_tendstoD)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1392
  with assms(2) have "eventually (\<lambda>x. False) net"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1393
    by (rule eventually_elim2) simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1394
  with assms(3) show "False"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1395
    by (simp add: eventually_False)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1396
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1397
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1398
text{* Need to prove closed(cball(x,e)) before deducing this as a corollary. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1399
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1400
lemma Lim_dist_ubound:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1401
  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
  1402
  shows "dist a l <= e"
44252
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1403
proof-
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1404
  have "dist a l \<in> {..e}"
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1405
  proof (rule Lim_in_closed_set)
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1406
    show "closed {..e}" by simp
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1407
    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
  1408
    show "\<not> trivial_limit net" by fact
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1409
    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
  1410
  qed
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1411
  thus ?thesis by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1412
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1413
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1414
lemma Lim_norm_ubound:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1415
  fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1416
  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
  1417
  shows "norm(l) <= e"
44252
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1418
proof-
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1419
  have "norm l \<in> {..e}"
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1420
  proof (rule Lim_in_closed_set)
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1421
    show "closed {..e}" by simp
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1422
    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
  1423
    show "\<not> trivial_limit net" by fact
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1424
    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
  1425
  qed
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1426
  thus ?thesis by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1427
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1428
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1429
lemma Lim_norm_lbound:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1430
  fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1431
  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
  1432
  shows "e \<le> norm l"
44252
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1433
proof-
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1434
  have "norm l \<in> {e..}"
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1435
  proof (rule Lim_in_closed_set)
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1436
    show "closed {e..}" by simp
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1437
    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
  1438
    show "\<not> trivial_limit net" by fact
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1439
    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
  1440
  qed
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1441
  thus ?thesis by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1442
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1443
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1444
text{* Uniqueness of the limit, when nontrivial. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1445
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1446
lemma tendsto_Lim:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1447
  fixes f :: "'a \<Rightarrow> 'b::t2_space"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1448
  shows "~(trivial_limit net) \<Longrightarrow> (f ---> l) net ==> Lim net f = l"
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41969
diff changeset
  1449
  unfolding Lim_def using tendsto_unique[of net f] by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1450
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1451
text{* Limit under bilinear function *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1452
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1453
lemma Lim_bilinear:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1454
  assumes "(f ---> l) net" and "(g ---> m) net" and "bounded_bilinear h"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1455
  shows "((\<lambda>x. h (f x) (g x)) ---> (h l m)) net"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1456
using `bounded_bilinear h` `(f ---> l) net` `(g ---> m) net`
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1457
by (rule bounded_bilinear.tendsto)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1458
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1459
text{* These are special for limits out of the same vector space. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1460
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1461
lemma Lim_within_id: "(id ---> a) (at a within s)"
45031
9583f2b56f85 add lemmas within_empty and tendsto_bot;
huffman
parents: 44909
diff changeset
  1462
  unfolding id_def by (rule tendsto_ident_at_within)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1463
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1464
lemma Lim_at_id: "(id ---> a) (at a)"
45031
9583f2b56f85 add lemmas within_empty and tendsto_bot;
huffman
parents: 44909
diff changeset
  1465
  unfolding id_def by (rule tendsto_ident_at)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1466
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1467
lemma Lim_at_zero:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1468
  fixes a :: "'a::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1469
  fixes l :: "'b::topological_space"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1470
  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
  1471
  using LIM_offset_zero LIM_offset_zero_cancel ..
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1472
44081
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44076
diff changeset
  1473
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
  1474
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1475
definition
44081
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44076
diff changeset
  1476
  netlimit :: "'a::t2_space filter \<Rightarrow> 'a" where
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1477
  "netlimit net = (SOME a. ((\<lambda>x. x) ---> a) net)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1478
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1479
lemma netlimit_within:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1480
  assumes "\<not> trivial_limit (at a within S)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1481
  shows "netlimit (at a within S) = a"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1482
unfolding netlimit_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1483
apply (rule some_equality)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1484
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
  1485
apply (rule tendsto_ident_at)
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41969
diff changeset
  1486
apply (erule tendsto_unique [OF assms])
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1487
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
  1488
apply (rule tendsto_ident_at)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1489
done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1490
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1491
lemma netlimit_at:
44072
5b970711fb39 class perfect_space inherits from topological_space;
huffman
parents: 43338
diff changeset
  1492
  fixes a :: "'a::{perfect_space,t2_space}"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1493
  shows "netlimit (at a) = a"
45031
9583f2b56f85 add lemmas within_empty and tendsto_bot;
huffman
parents: 44909
diff changeset
  1494
  using netlimit_within [of a UNIV] by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1495
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1496
lemma lim_within_interior:
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1497
  "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
  1498
  by (simp add: at_within_interior)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1499
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1500
lemma netlimit_within_interior:
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1501
  fixes x :: "'a::{t2_space,perfect_space}"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1502
  assumes "x \<in> interior S"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1503
  shows "netlimit (at x within S) = x"
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1504
using assms by (simp add: at_within_interior netlimit_at)
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1505
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1506
text{* Transformation of limit. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1507
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1508
lemma Lim_transform:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1509
  fixes f g :: "'a::type \<Rightarrow> 'b::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1510
  assumes "((\<lambda>x. f x - g x) ---> 0) net" "(f ---> l) net"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1511
  shows "(g ---> l) net"
44252
10362a07eb7c Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents: 44250
diff changeset
  1512
  using tendsto_diff [OF assms(2) assms(1)] by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1513
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1514
lemma Lim_transform_eventually:
36667
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1515
  "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
  1516
  apply (rule topological_tendstoI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1517
  apply (drule (2) topological_tendstoD)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1518
  apply (erule (1) eventually_elim2, simp)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1519
  done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1520
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1521
lemma Lim_transform_within:
36667
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1522
  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
  1523
  and "(f ---> l) (at x within S)"
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1524
  shows "(g ---> l) (at x within S)"
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1525
proof (rule Lim_transform_eventually)
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1526
  show "eventually (\<lambda>x. f x = g x) (at x within S)"
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1527
    unfolding eventually_within
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1528
    using assms(1,2) by auto
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1529
  show "(f ---> l) (at x within S)" by fact
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1530
qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1531
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1532
lemma Lim_transform_at:
36667
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1533
  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
  1534
  and "(f ---> l) (at x)"
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1535
  shows "(g ---> l) (at x)"
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1536
proof (rule Lim_transform_eventually)
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1537
  show "eventually (\<lambda>x. f x = g x) (at x)"
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1538
    unfolding eventually_at
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1539
    using assms(1,2) by auto
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1540
  show "(f ---> l) (at x)" by fact
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1541
qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1542
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1543
text{* Common case assuming being away from some crucial point like 0. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1544
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1545
lemma Lim_transform_away_within:
36669
c90c8a3ae1f7 generalize some lemmas to class t1_space
huffman
parents: 36668
diff changeset
  1546
  fixes a b :: "'a::t1_space"
36667
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1547
  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
  1548
  and "(f ---> l) (at a within S)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1549
  shows "(g ---> l) (at a within S)"
36669
c90c8a3ae1f7 generalize some lemmas to class t1_space
huffman
parents: 36668
diff changeset
  1550
proof (rule Lim_transform_eventually)
c90c8a3ae1f7 generalize some lemmas to class t1_space
huffman
parents: 36668
diff changeset
  1551
  show "(f ---> l) (at a within S)" by fact
c90c8a3ae1f7 generalize some lemmas to class t1_space
huffman
parents: 36668
diff changeset
  1552
  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
  1553
    unfolding Limits.eventually_within eventually_at_topological
c90c8a3ae1f7 generalize some lemmas to class t1_space
huffman
parents: 36668
diff changeset
  1554
    by (rule exI [where x="- {b}"], simp add: open_Compl assms)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1555
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1556
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1557
lemma Lim_transform_away_at:
36669
c90c8a3ae1f7 generalize some lemmas to class t1_space
huffman
parents: 36668
diff changeset
  1558
  fixes a b :: "'a::t1_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1559
  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
  1560
  and fl: "(f ---> l) (at a)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1561
  shows "(g ---> l) (at a)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1562
  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
  1563
  by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1564
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1565
text{* Alternatively, within an open set. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1566
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1567
lemma Lim_transform_within_open:
36667
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1568
  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
  1569
  and "(f ---> l) (at a)"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1570
  shows "(g ---> l) (at a)"
36667
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1571
proof (rule Lim_transform_eventually)
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1572
  show "eventually (\<lambda>x. f x = g x) (at a)"
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1573
    unfolding eventually_at_topological
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1574
    using assms(1,2,3) by auto
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1575
  show "(f ---> l) (at a)" by fact
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1576
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1577
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1578
text{* A congruence rule allowing us to transform limits assuming not at point. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1579
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1580
(* FIXME: Only one congruence rule for tendsto can be used at a time! *)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1581
36362
06475a1547cb fix lots of looping simp calls and other warnings
huffman
parents: 36360
diff changeset
  1582
lemma Lim_cong_within(*[cong add]*):
43338
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1583
  assumes "a = b" "x = y" "S = T"
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1584
  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
  1585
  shows "(f ---> x) (at a within S) \<longleftrightarrow> (g ---> y) (at b within T)"
36667
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1586
  unfolding tendsto_def Limits.eventually_within eventually_at_topological
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1587
  using assms by simp
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1588
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1589
lemma Lim_cong_at(*[cong add]*):
43338
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1590
  assumes "a = b" "x = y"
36667
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1591
  assumes "\<And>x. x \<noteq> a \<Longrightarrow> f x = g x"
43338
a150d16bf77c lemmas about right derivative and limits
hoelzl
parents: 42165
diff changeset
  1592
  shows "((\<lambda>x. f x) ---> x) (at a) \<longleftrightarrow> ((g ---> y) (at a))"
36667
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1593
  unfolding tendsto_def eventually_at_topological
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1594
  using assms by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1595
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1596
text{* Useful lemmas on closure and set of possible sequential limits.*}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1597
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1598
lemma closure_sequential:
36667
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  1599
  fixes l :: "'a::metric_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1600
  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
  1601
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1602
  assume "?lhs" moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1603
  { assume "l \<in> S"
44125
230a8665c919 mark some redundant theorems as legacy
huffman
parents: 44122
diff changeset
  1604
    hence "?rhs" using tendsto_const[of l sequentially] by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1605
  } moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1606
  { assume "l islimpt S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1607
    hence "?rhs" unfolding islimpt_sequential by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1608
  } ultimately
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1609
  show "?rhs" unfolding closure_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1610
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1611
  assume "?rhs"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1612
  thus "?lhs" unfolding closure_def unfolding islimpt_sequential by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1613
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1614
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1615
lemma closed_sequential_limits:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1616
  fixes S :: "'a::metric_space set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1617
  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
  1618
  unfolding closed_limpt
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1619
  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
  1620
  by metis
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1621
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1622
lemma closure_approachable:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1623
  fixes S :: "'a::metric_space set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1624
  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
  1625
  apply (auto simp add: closure_def islimpt_approachable)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1626
  by (metis dist_self)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1627
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1628
lemma closed_approachable:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1629
  fixes S :: "'a::metric_space set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1630
  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
  1631
  by (metis closure_closed closure_approachable)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1632
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1633
subsection {* Infimum Distance *}
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1634
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1635
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
  1636
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1637
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
  1638
  by (simp add: infdist_def)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1639
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1640
lemma infdist_nonneg:
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1641
  shows "0 \<le> infdist x A"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1642
  using assms by (auto simp add: infdist_def)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1643
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1644
lemma infdist_le:
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1645
  assumes "a \<in> A"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1646
  assumes "d = dist x a"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1647
  shows "infdist x A \<le> d"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1648
  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
  1649
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1650
lemma infdist_zero[simp]:
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1651
  assumes "a \<in> A" shows "infdist a A = 0"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1652
proof -
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1653
  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
  1654
  with infdist_nonneg[of a A] assms show "infdist a A = 0" by auto
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1655
qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1656
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1657
lemma infdist_triangle:
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1658
  shows "infdist x A \<le> infdist y A + dist x y"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1659
proof cases
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1660
  assume "A = {}" thus ?thesis by (simp add: infdist_def)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1661
next
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1662
  assume "A \<noteq> {}" then obtain a where "a \<in> A" by auto
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1663
  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
  1664
  proof
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1665
    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
  1666
    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
  1667
    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
  1668
    show "infdist x A \<le> d"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1669
      unfolding infdist_notempty[OF `A \<noteq> {}`]
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1670
    proof (rule Inf_lower2)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1671
      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
  1672
      show "dist x a \<le> d" unfolding d by (rule dist_triangle)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1673
      fix d assume "d \<in> {dist x a |a. a \<in> A}"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1674
      then obtain a where "a \<in> A" "d = dist x a" by auto
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1675
      thus "infdist x A \<le> d" by (rule infdist_le)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1676
    qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1677
  qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1678
  also have "\<dots> = dist x y + infdist y A"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1679
  proof (rule Inf_eq, safe)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1680
    fix a assume "a \<in> A"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1681
    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
  1682
  next
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1683
    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
  1684
    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
  1685
      by (intro Inf_greatest) (auto simp: field_simps)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1686
    thus "i \<le> dist x y + infdist y A" by simp
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1687
  qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1688
  finally show ?thesis by simp
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1689
qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1690
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1691
lemma
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1692
  in_closure_iff_infdist_zero:
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1693
  assumes "A \<noteq> {}"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1694
  shows "x \<in> closure A \<longleftrightarrow> infdist x A = 0"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1695
proof
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1696
  assume "x \<in> closure A"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1697
  show "infdist x A = 0"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1698
  proof (rule ccontr)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1699
    assume "infdist x A \<noteq> 0"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1700
    with infdist_nonneg[of x A] have "infdist x A > 0" by auto
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1701
    hence "ball x (infdist x A) \<inter> closure A = {}" apply auto
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1702
      by (metis `0 < infdist x A` `x \<in> closure A` closure_approachable dist_commute
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1703
        eucl_less_not_refl euclidean_trans(2) infdist_le)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1704
    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
  1705
    thus False using `x \<in> closure A` by simp
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1706
  qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1707
next
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1708
  assume x: "infdist x A = 0"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1709
  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
  1710
  show "x \<in> closure A" unfolding closure_approachable
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1711
  proof (safe, rule ccontr)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1712
    fix e::real assume "0 < e"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1713
    assume "\<not> (\<exists>y\<in>A. dist y x < e)"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1714
    hence "infdist x A \<ge> e" using `a \<in> A`
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1715
      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
  1716
      by (force simp: dist_commute)
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1717
    with x `0 < e` show False by auto
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1718
  qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1719
qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1720
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1721
lemma
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1722
  in_closed_iff_infdist_zero:
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1723
  assumes "closed A" "A \<noteq> {}"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1724
  shows "x \<in> A \<longleftrightarrow> infdist x A = 0"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1725
proof -
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1726
  have "x \<in> closure A \<longleftrightarrow> infdist x A = 0"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1727
    by (rule in_closure_iff_infdist_zero) fact
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1728
  with assms show ?thesis by simp
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1729
qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1730
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1731
lemma tendsto_infdist [tendsto_intros]:
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1732
  assumes f: "(f ---> l) F"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1733
  shows "((\<lambda>x. infdist (f x) A) ---> infdist l A) F"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1734
proof (rule tendstoI)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1735
  fix e ::real assume "0 < e"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1736
  from tendstoD[OF f this]
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1737
  show "eventually (\<lambda>x. dist (infdist (f x) A) (infdist l A) < e) F"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1738
  proof (eventually_elim)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1739
    fix x
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1740
    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
  1741
    have "dist (infdist (f x) A) (infdist l A) \<le> dist (f x) l"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1742
      by (simp add: dist_commute dist_real_def)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1743
    also assume "dist (f x) l < e"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1744
    finally show "dist (infdist (f x) A) (infdist l A) < e" .
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1745
  qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1746
qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  1747
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1748
text{* Some other lemmas about sequences. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1749
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  1750
lemma sequentially_offset:
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  1751
  assumes "eventually (\<lambda>i. P i) sequentially"
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  1752
  shows "eventually (\<lambda>i. P (i + k)) sequentially"
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  1753
  using assms unfolding eventually_sequentially by (metis trans_le_add1)
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  1754
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1755
lemma seq_offset:
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  1756
  assumes "(f ---> l) sequentially"
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  1757
  shows "((\<lambda>i. f (i + k)) ---> l) sequentially"
44584
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  1758
  using assms by (rule LIMSEQ_ignore_initial_segment) (* FIXME: redundant *)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1759
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1760
lemma seq_offset_neg:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1761
  "(f ---> l) sequentially ==> ((\<lambda>i. f(i - k)) ---> l) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1762
  apply (rule topological_tendstoI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1763
  apply (drule (2) topological_tendstoD)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1764
  apply (simp only: eventually_sequentially)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1765
  apply (subgoal_tac "\<And>N k (n::nat). N + k <= n ==> N <= n - k")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1766
  apply metis
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1767
  by arith
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1768
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1769
lemma seq_offset_rev:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1770
  "((\<lambda>i. f(i + k)) ---> l) sequentially ==> (f ---> l) sequentially"
44584
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  1771
  by (rule LIMSEQ_offset) (* FIXME: redundant *)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1772
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1773
lemma seq_harmonic: "((\<lambda>n. inverse (real n)) ---> 0) sequentially"
44584
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  1774
  using LIMSEQ_inverse_real_of_nat by (rule LIMSEQ_imp_Suc)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1775
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  1776
subsection {* More properties of closed balls *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1777
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1778
lemma closed_cball: "closed (cball x e)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1779
unfolding cball_def closed_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1780
unfolding Collect_neg_eq [symmetric] not_le
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1781
apply (clarsimp simp add: open_dist, rename_tac y)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1782
apply (rule_tac x="dist x y - e" in exI, clarsimp)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1783
apply (rename_tac x')
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1784
apply (cut_tac x=x and y=x' and z=y in dist_triangle)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1785
apply simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1786
done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1787
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1788
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
  1789
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1790
  { 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
  1791
    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
  1792
  } moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1793
  { 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
  1794
    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
  1795
  } ultimately
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1796
  show ?thesis unfolding open_contains_ball by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1797
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1798
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1799
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
  1800
  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
  1801
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1802
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
  1803
  apply (simp add: interior_def, safe)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1804
  apply (force simp add: open_contains_cball)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1805
  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
  1806
  apply (simp add: subset_trans [OF ball_subset_cball])
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1807
  done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1808
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1809
lemma islimpt_ball:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1810
  fixes x y :: "'a::{real_normed_vector,perfect_space}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1811
  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
  1812
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1813
  assume "?lhs"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1814
  { assume "e \<le> 0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1815
    hence *:"ball x e = {}" using ball_eq_empty[of x e] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1816
    have False using `?lhs` unfolding * using islimpt_EMPTY[of y] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1817
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1818
  hence "e > 0" by (metis not_less)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1819
  moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1820
  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
  1821
  ultimately show "?rhs" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1822
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1823
  assume "?rhs" hence "e>0"  by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1824
  { fix d::real assume "d>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1825
    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
  1826
    proof(cases "d \<le> dist x y")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1827
      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
  1828
      proof(cases "x=y")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1829
        case True hence False using `d \<le> dist x y` `d>0` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1830
        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
  1831
      next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1832
        case False
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1833
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1834
        have "dist x (y - (d / (2 * dist y x)) *\<^sub>R (y - x))
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1835
              = norm (x - y + (d / (2 * norm (y - x))) *\<^sub>R (y - x))"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1836
          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
  1837
        also have "\<dots> = \<bar>- 1 + d / (2 * norm (x - y))\<bar> * norm (x - y)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1838
          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
  1839
          unfolding scaleR_minus_left scaleR_one
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1840
          by (auto simp add: norm_minus_commute)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1841
        also have "\<dots> = \<bar>- norm (x - y) + d / 2\<bar>"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1842
          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
  1843
          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
  1844
        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
  1845
        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
  1846
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1847
        moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1848
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1849
        have "(d / (2*dist y x)) *\<^sub>R (y - x) \<noteq> 0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1850
          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
  1851
        moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1852
        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
  1853
          using `d>0` `x\<noteq>y`[unfolded dist_nz] dist_commute[of x y]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1854
          unfolding dist_norm by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1855
        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
  1856
      qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1857
    next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1858
      case False hence "d > dist x y" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1859
      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
  1860
      proof(cases "x=y")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1861
        case True
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1862
        obtain z where **: "z \<noteq> y" "dist z y < min e d"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1863
          using perfect_choose_dist[of "min e d" y]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1864
          using `d > 0` `e>0` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1865
        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
  1866
          unfolding `x = y`
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1867
          using `z \<noteq> y` **
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1868
          by (rule_tac x=z in bexI, auto simp add: dist_commute)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1869
      next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1870
        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
  1871
          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
  1872
      qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1873
    qed  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1874
  thus "?lhs" unfolding mem_cball islimpt_approachable mem_ball by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1875
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1876
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1877
lemma closure_ball_lemma:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1878
  fixes x y :: "'a::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1879
  assumes "x \<noteq> y" shows "y islimpt ball x (dist x y)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1880
proof (rule islimptI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1881
  fix T assume "y \<in> T" "open T"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1882
  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
  1883
    unfolding open_dist by fast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1884
  (* choose point between x and y, within distance r of y. *)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1885
  def k \<equiv> "min 1 (r / (2 * dist x y))"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1886
  def z \<equiv> "y + scaleR k (x - y)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1887
  have z_def2: "z = x + scaleR (1 - k) (y - x)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1888
    unfolding z_def by (simp add: algebra_simps)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1889
  have "dist z y < r"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1890
    unfolding z_def k_def using `0 < r`
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1891
    by (simp add: dist_norm min_def)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1892
  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
  1893
  have "dist x z < dist x y"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1894
    unfolding z_def2 dist_norm
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1895
    apply (simp add: norm_minus_commute)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1896
    apply (simp only: dist_norm [symmetric])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1897
    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
  1898
    apply (rule mult_strict_right_mono)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1899
    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
  1900
    apply (simp add: zero_less_dist_iff `x \<noteq> y`)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1901
    done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1902
  hence "z \<in> ball x (dist x y)" by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1903
  have "z \<noteq> y"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1904
    unfolding z_def k_def using `x \<noteq> y` `0 < r`
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1905
    by (simp add: min_def)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1906
  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
  1907
    using `z \<in> ball x (dist x y)` `z \<in> T` `z \<noteq> y`
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1908
    by fast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1909
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1910
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1911
lemma closure_ball:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1912
  fixes x :: "'a::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1913
  shows "0 < e \<Longrightarrow> closure (ball x e) = cball x e"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1914
apply (rule equalityI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1915
apply (rule closure_minimal)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1916
apply (rule ball_subset_cball)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1917
apply (rule closed_cball)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1918
apply (rule subsetI, rename_tac y)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1919
apply (simp add: le_less [where 'a=real])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1920
apply (erule disjE)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1921
apply (rule subsetD [OF closure_subset], simp)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1922
apply (simp add: closure_def)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1923
apply clarify
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1924
apply (rule closure_ball_lemma)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1925
apply (simp add: zero_less_dist_iff)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1926
done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1927
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1928
(* In a trivial vector space, this fails for e = 0. *)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1929
lemma interior_cball:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1930
  fixes x :: "'a::{real_normed_vector, perfect_space}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1931
  shows "interior (cball x e) = ball x e"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1932
proof(cases "e\<ge>0")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1933
  case False note cs = this
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1934
  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
  1935
  { fix y assume "y \<in> cball x e"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1936
    hence False unfolding mem_cball using dist_nz[of x y] cs by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1937
  hence "cball x e = {}" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1938
  hence "interior (cball x e) = {}" using interior_empty by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1939
  ultimately show ?thesis by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1940
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1941
  case True note cs = this
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1942
  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
  1943
  { 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
  1944
    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
  1945
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1946
    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
  1947
      using perfect_choose_dist [of d] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1948
    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
  1949
    hence xa_cball:"xa \<in> cball x e" using as(1) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1950
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1951
    hence "y \<in> ball x e" proof(cases "x = y")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1952
      case True
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1953
      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
  1954
      thus "y \<in> ball x e" using `x = y ` by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1955
    next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1956
      case False
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1957
      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
  1958
        using `d>0` norm_ge_zero[of "y - x"] `x \<noteq> y` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1959
      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
  1960
      have "y - x \<noteq> 0" using `x \<noteq> y` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1961
      hence **:"d / (2 * norm (y - x)) > 0" unfolding zero_less_norm_iff[THEN sym]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1962
        using `d>0` divide_pos_pos[of d "2*norm (y - x)"] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1963
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1964
      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
  1965
        by (auto simp add: dist_norm algebra_simps)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1966
      also have "\<dots> = norm ((1 + d / (2 * norm (y - x))) *\<^sub>R (y - x))"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1967
        by (auto simp add: algebra_simps)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1968
      also have "\<dots> = \<bar>1 + d / (2 * norm (y - x))\<bar> * norm (y - x)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1969
        using ** by auto
49962
a8cc904a6820 Renamed {left,right}_distrib to distrib_{right,left}.
webertj
parents: 49834
diff changeset
  1970
      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
  1971
      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
  1972
      thus "y \<in> ball x e" unfolding mem_ball using `d>0` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1973
    qed  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1974
  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
  1975
  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
  1976
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1977
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1978
lemma frontier_ball:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1979
  fixes a :: "'a::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1980
  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
  1981
  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
  1982
  apply (simp add: set_eq_iff)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1983
  by arith
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1984
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1985
lemma frontier_cball:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1986
  fixes a :: "'a::{real_normed_vector, perfect_space}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1987
  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
  1988
  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
  1989
  apply (simp add: set_eq_iff)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1990
  by arith
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1991
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1992
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
  1993
  apply (simp add: set_eq_iff not_le)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1994
  by (metis zero_le_dist dist_self order_less_le_trans)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1995
lemma cball_empty: "e < 0 ==> cball x e = {}" by (simp add: cball_eq_empty)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1996
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1997
lemma cball_eq_sing:
44072
5b970711fb39 class perfect_space inherits from topological_space;
huffman
parents: 43338
diff changeset
  1998
  fixes x :: "'a::{metric_space,perfect_space}"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1999
  shows "(cball x e = {x}) \<longleftrightarrow> e = 0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2000
proof (rule linorder_cases)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2001
  assume e: "0 < e"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2002
  obtain a where "a \<noteq> x" "dist a x < e"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2003
    using perfect_choose_dist [OF e] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2004
  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
  2005
  with e show ?thesis by (auto simp add: set_eq_iff)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2006
qed auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2007
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2008
lemma cball_sing:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2009
  fixes x :: "'a::metric_space"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2010
  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
  2011
  by (auto simp add: set_eq_iff)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2012
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  2013
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  2014
subsection {* Boundedness *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2015
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2016
  (* 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
  2017
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
  2018
  bounded :: "'a set \<Rightarrow> bool" where
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2019
  "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
  2020
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2021
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
  2022
unfolding bounded_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2023
apply safe
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2024
apply (rule_tac x="dist a x + e" in exI, clarify)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2025
apply (drule (1) bspec)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2026
apply (erule order_trans [OF dist_triangle add_left_mono])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2027
apply auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2028
done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2029
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2030
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
  2031
unfolding bounded_any_center [where a=0]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2032
by (simp add: dist_norm)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2033
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50094
diff changeset
  2034
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
  2035
  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
  2036
  using assms by auto
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50094
diff changeset
  2037
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2038
lemma bounded_empty[simp]: "bounded {}" by (simp add: bounded_def)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2039
lemma bounded_subset: "bounded T \<Longrightarrow> S \<subseteq> T ==> bounded S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2040
  by (metis bounded_def subset_eq)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2041
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2042
lemma bounded_interior[intro]: "bounded S ==> bounded(interior S)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2043
  by (metis bounded_subset interior_subset)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2044
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2045
lemma bounded_closure[intro]: assumes "bounded S" shows "bounded(closure S)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2046
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2047
  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
  2048
  { fix y assume "y \<in> closure S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2049
    then obtain f where f: "\<forall>n. f n \<in> S"  "(f ---> y) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2050
      unfolding closure_sequential by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2051
    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
  2052
    hence "eventually (\<lambda>n. dist x (f n) \<le> a) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2053
      by (rule eventually_mono, simp add: f(1))
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2054
    have "dist x y \<le> a"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2055
      apply (rule Lim_dist_ubound [of sequentially f])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2056
      apply (rule trivial_limit_sequentially)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2057
      apply (rule f(2))
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2058
      apply fact
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2059
      done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2060
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2061
  thus ?thesis unfolding bounded_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2062
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2063
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2064
lemma bounded_cball[simp,intro]: "bounded (cball x e)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2065
  apply (simp add: bounded_def)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2066
  apply (rule_tac x=x in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2067
  apply (rule_tac x=e in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2068
  apply auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2069
  done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2070
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2071
lemma bounded_ball[simp,intro]: "bounded(ball x e)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2072
  by (metis ball_subset_cball bounded_cball bounded_subset)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2073
36362
06475a1547cb fix lots of looping simp calls and other warnings
huffman
parents: 36360
diff changeset
  2074
lemma finite_imp_bounded[intro]:
06475a1547cb fix lots of looping simp calls and other warnings
huffman
parents: 36360
diff changeset
  2075
  fixes S :: "'a::metric_space set" assumes "finite S" shows "bounded S"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2076
proof-
36362
06475a1547cb fix lots of looping simp calls and other warnings
huffman
parents: 36360
diff changeset
  2077
  { fix a and F :: "'a set" assume as:"bounded F"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2078
    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
  2079
    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
  2080
    hence "bounded (insert a F)" unfolding bounded_def by (intro exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2081
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2082
  thus ?thesis using finite_induct[of S bounded]  using bounded_empty assms 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 bounded_Un[simp]: "bounded (S \<union> T) \<longleftrightarrow> bounded S \<and> bounded T"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2086
  apply (auto simp add: bounded_def)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2087
  apply (rename_tac x y r s)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2088
  apply (rule_tac x=x in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2089
  apply (rule_tac x="max r (dist x y + s)" in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2090
  apply (rule ballI, rename_tac z, safe)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2091
  apply (drule (1) bspec, simp)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2092
  apply (drule (1) bspec)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2093
  apply (rule min_max.le_supI2)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2094
  apply (erule order_trans [OF dist_triangle add_left_mono])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2095
  done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2096
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2097
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
  2098
  by (induct rule: finite_induct[of F], auto)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2099
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2100
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
  2101
  apply (simp add: bounded_iff)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2102
  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
  2103
  by metis arith
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2104
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2105
lemma bounded_Int[intro]: "bounded S \<or> bounded T \<Longrightarrow> bounded (S \<inter> T)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2106
  by (metis Int_lower1 Int_lower2 bounded_subset)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2107
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2108
lemma bounded_diff[intro]: "bounded S ==> bounded (S - T)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2109
apply (metis Diff_subset bounded_subset)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2110
done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2111
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2112
lemma bounded_insert[intro]:"bounded(insert x S) \<longleftrightarrow> bounded S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2113
  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
  2114
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2115
lemma not_bounded_UNIV[simp, intro]:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2116
  "\<not> bounded (UNIV :: 'a::{real_normed_vector, perfect_space} set)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2117
proof(auto simp add: bounded_pos not_le)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2118
  obtain x :: 'a where "x \<noteq> 0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2119
    using perfect_choose_dist [OF zero_less_one] by fast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2120
  fix b::real  assume b: "b >0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2121
  have b1: "b +1 \<ge> 0" using b by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2122
  with `x \<noteq> 0` have "b < norm (scaleR (b + 1) (sgn x))"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2123
    by (simp add: norm_sgn)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2124
  then show "\<exists>x::'a. b < norm x" ..
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2125
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2126
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2127
lemma bounded_linear_image:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2128
  assumes "bounded S" "bounded_linear f"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2129
  shows "bounded(f ` S)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2130
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2131
  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
  2132
  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
  2133
  { fix x assume "x\<in>S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2134
    hence "norm x \<le> b" using b by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2135
    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
  2136
      by (metis B(1) B(2) order_trans mult_le_cancel_left_pos)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2137
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2138
  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
  2139
    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
  2140
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2141
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2142
lemma bounded_scaling:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2143
  fixes S :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2144
  shows "bounded S \<Longrightarrow> bounded ((\<lambda>x. c *\<^sub>R x) ` S)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2145
  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
  2146
  apply (rule bounded_linear_scaleR_right)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2147
  done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2148
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2149
lemma bounded_translation:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2150
  fixes S :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2151
  assumes "bounded S" shows "bounded ((\<lambda>x. a + x) ` S)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2152
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2153
  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
  2154
  { fix x assume "x\<in>S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2155
    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
  2156
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2157
  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
  2158
    by (auto intro!: exI[of _ "b + norm a"])
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2159
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2160
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2161
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2162
text{* Some theorems on sups and infs using the notion "bounded". *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2163
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2164
lemma bounded_real:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2165
  fixes S :: "real set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2166
  shows "bounded S \<longleftrightarrow>  (\<exists>a. \<forall>x\<in>S. abs x <= a)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2167
  by (simp add: bounded_iff)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2168
33270
paulson
parents: 33175
diff changeset
  2169
lemma bounded_has_Sup:
paulson
parents: 33175
diff changeset
  2170
  fixes S :: "real set"
paulson
parents: 33175
diff changeset
  2171
  assumes "bounded S" "S \<noteq> {}"
paulson
parents: 33175
diff changeset
  2172
  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
  2173
proof
paulson
parents: 33175
diff changeset
  2174
  fix x assume "x\<in>S"
paulson
parents: 33175
diff changeset
  2175
  thus "x \<le> Sup S"
paulson
parents: 33175
diff changeset
  2176
    by (metis SupInf.Sup_upper abs_le_D1 assms(1) bounded_real)
paulson
parents: 33175
diff changeset
  2177
next
paulson
parents: 33175
diff changeset
  2178
  show "\<forall>b. (\<forall>x\<in>S. x \<le> b) \<longrightarrow> Sup S \<le> b" using assms
paulson
parents: 33175
diff changeset
  2179
    by (metis SupInf.Sup_least)
paulson
parents: 33175
diff changeset
  2180
qed
paulson
parents: 33175
diff changeset
  2181
paulson
parents: 33175
diff changeset
  2182
lemma Sup_insert:
paulson
parents: 33175
diff changeset
  2183
  fixes S :: "real set"
paulson
parents: 33175
diff changeset
  2184
  shows "bounded S ==> Sup(insert x S) = (if S = {} then x else max x (Sup S))" 
paulson
parents: 33175
diff changeset
  2185
by auto (metis Int_absorb Sup_insert_nonempty assms bounded_has_Sup(1) disjoint_iff_not_equal) 
paulson
parents: 33175
diff changeset
  2186
paulson
parents: 33175
diff changeset
  2187
lemma Sup_insert_finite:
paulson
parents: 33175
diff changeset
  2188
  fixes S :: "real set"
paulson
parents: 33175
diff changeset
  2189
  shows "finite S \<Longrightarrow> Sup(insert x S) = (if S = {} then x else max x (Sup S))"
paulson
parents: 33175
diff changeset
  2190
  apply (rule Sup_insert)
paulson
parents: 33175
diff changeset
  2191
  apply (rule finite_imp_bounded)
paulson
parents: 33175
diff changeset
  2192
  by simp
paulson
parents: 33175
diff changeset
  2193
paulson
parents: 33175
diff changeset
  2194
lemma bounded_has_Inf:
paulson
parents: 33175
diff changeset
  2195
  fixes S :: "real set"
paulson
parents: 33175
diff changeset
  2196
  assumes "bounded S"  "S \<noteq> {}"
paulson
parents: 33175
diff changeset
  2197
  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
  2198
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2199
  fix x assume "x\<in>S"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2200
  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
  2201
  thus "x \<ge> Inf S" using `x\<in>S`
paulson
parents: 33175
diff changeset
  2202
    by (metis Inf_lower_EX abs_le_D2 minus_le_iff)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2203
next
33270
paulson
parents: 33175
diff changeset
  2204
  show "\<forall>b. (\<forall>x\<in>S. x >= b) \<longrightarrow> Inf S \<ge> b" using assms
paulson
parents: 33175
diff changeset
  2205
    by (metis SupInf.Inf_greatest)
paulson
parents: 33175
diff changeset
  2206
qed
paulson
parents: 33175
diff changeset
  2207
paulson
parents: 33175
diff changeset
  2208
lemma Inf_insert:
paulson
parents: 33175
diff changeset
  2209
  fixes S :: "real set"
paulson
parents: 33175
diff changeset
  2210
  shows "bounded S ==> Inf(insert x S) = (if S = {} then x else min x (Inf S))" 
paulson
parents: 33175
diff changeset
  2211
by auto (metis Int_absorb Inf_insert_nonempty bounded_has_Inf(1) disjoint_iff_not_equal) 
paulson
parents: 33175
diff changeset
  2212
lemma Inf_insert_finite:
paulson
parents: 33175
diff changeset
  2213
  fixes S :: "real set"
paulson
parents: 33175
diff changeset
  2214
  shows "finite S ==> Inf(insert x S) = (if S = {} then x else min x (Inf S))"
paulson
parents: 33175
diff changeset
  2215
  by (rule Inf_insert, rule finite_imp_bounded, simp)
paulson
parents: 33175
diff changeset
  2216
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2217
(* TODO: Move this to RComplete.thy -- would need to include Glb into RComplete *)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2218
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
  2219
  apply (frule isGlb_isLb)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2220
  apply (frule_tac x = y in isGlb_isLb)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2221
  apply (blast intro!: order_antisym dest!: isGlb_le_isLb)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2222
  done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2223
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  2224
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  2225
subsection {* Equivalent versions of compactness *}
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  2226
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  2227
subsubsection{* Sequential compactness *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2228
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2229
definition
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2230
  compact :: "'a::metric_space set \<Rightarrow> bool" where (* TODO: generalize *)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2231
  "compact S \<longleftrightarrow>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2232
   (\<forall>f. (\<forall>n. f n \<in> S) \<longrightarrow>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2233
       (\<exists>l\<in>S. \<exists>r. subseq r \<and> ((f o r) ---> l) sequentially))"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2234
44075
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  2235
lemma compactI:
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  2236
  assumes "\<And>f. \<forall>n. f n \<in> S \<Longrightarrow> \<exists>l\<in>S. \<exists>r. subseq r \<and> ((f o r) ---> l) sequentially"
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  2237
  shows "compact S"
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  2238
  unfolding compact_def using assms by fast
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  2239
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  2240
lemma compactE:
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  2241
  assumes "compact S" "\<forall>n. f n \<in> S"
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  2242
  obtains l r where "l \<in> S" "subseq r" "((f \<circ> r) ---> l) sequentially"
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  2243
  using assms unfolding compact_def by fast
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  2244
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2245
text {*
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2246
  A metric space (or topological vector space) is said to have the
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2247
  Heine-Borel property if every closed and bounded subset is compact.
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2248
*}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2249
44207
ea99698c2070 locale-ize some definitions, so perfect_space and heine_borel can inherit from the proper superclasses
huffman
parents: 44170
diff changeset
  2250
class heine_borel = metric_space +
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2251
  assumes bounded_imp_convergent_subsequence:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2252
    "bounded s \<Longrightarrow> \<forall>n. f n \<in> s
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2253
      \<Longrightarrow> \<exists>l r. subseq r \<and> ((f \<circ> r) ---> l) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2254
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2255
lemma bounded_closed_imp_compact:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2256
  fixes s::"'a::heine_borel set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2257
  assumes "bounded s" and "closed s" shows "compact s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2258
proof (unfold compact_def, clarify)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2259
  fix f :: "nat \<Rightarrow> 'a" assume f: "\<forall>n. f n \<in> s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2260
  obtain l r where r: "subseq r" and l: "((f \<circ> r) ---> l) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2261
    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
  2262
  from f have fr: "\<forall>n. (f \<circ> r) n \<in> s" by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2263
  have "l \<in> s" using `closed s` fr l
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2264
    unfolding closed_sequential_limits by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2265
  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
  2266
    using `l \<in> s` r l by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2267
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2268
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2269
lemma subseq_bigger: assumes "subseq r" shows "n \<le> r n"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2270
proof(induct n)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2271
  show "0 \<le> r 0" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2272
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2273
  fix n assume "n \<le> r n"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2274
  moreover have "r n < r (Suc n)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2275
    using assms [unfolded subseq_def] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2276
  ultimately show "Suc n \<le> r (Suc n)" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2277
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2278
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2279
lemma eventually_subseq:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2280
  assumes r: "subseq r"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2281
  shows "eventually P sequentially \<Longrightarrow> eventually (\<lambda>n. P (r n)) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2282
unfolding eventually_sequentially
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2283
by (metis subseq_bigger [OF r] le_trans)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2284
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2285
lemma lim_subseq:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2286
  "subseq r \<Longrightarrow> (s ---> l) sequentially \<Longrightarrow> ((s o r) ---> l) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2287
unfolding tendsto_def eventually_sequentially o_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2288
by (metis subseq_bigger le_trans)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2289
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2290
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
  2291
  unfolding Ex1_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2292
  apply (rule_tac x="nat_rec e f" in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2293
  apply (rule conjI)+
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2294
apply (rule def_nat_rec_0, simp)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2295
apply (rule allI, rule def_nat_rec_Suc, simp)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2296
apply (rule allI, rule impI, rule ext)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2297
apply (erule conjE)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2298
apply (induct_tac x)
36362
06475a1547cb fix lots of looping simp calls and other warnings
huffman
parents: 36360
diff changeset
  2299
apply simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2300
apply (erule_tac x="n" in allE)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2301
apply (simp)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2302
done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2303
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2304
lemma convergent_bounded_increasing: fixes s ::"nat\<Rightarrow>real"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2305
  assumes "incseq s" and "\<forall>n. abs(s n) \<le> b"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2306
  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
  2307
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2308
  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
  2309
  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
  2310
  { 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
  2311
    { fix n::nat
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2312
      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
  2313
      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
  2314
      with n have "s N \<le> t - e" using `e>0` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2315
      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
  2316
    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
  2317
    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
  2318
  thus ?thesis by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2319
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2320
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2321
lemma convergent_bounded_monotone: fixes s::"nat \<Rightarrow> real"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2322
  assumes "\<forall>n. abs(s n) \<le> b" and "monoseq s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2323
  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
  2324
  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
  2325
  unfolding monoseq_def incseq_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2326
  apply auto unfolding minus_add_distrib[THEN sym, unfolded diff_minus[THEN sym]]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2327
  unfolding abs_minus_cancel by(rule_tac x="-l" in exI)auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2328
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  2329
(* 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
  2330
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
  2331
  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
  2332
  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
  2333
proof-
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  2334
  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
  2335
  { 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
  2336
    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
  2337
      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
  2338
      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
  2339
  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
  2340
  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
  2341
    unfolding monoseq_def by auto
44907
93943da0a010 remove redundant lemma Lim_sequentially in favor of lemma LIMSEQ_def
huffman
parents: 44905
diff changeset
  2342
  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
  2343
    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
  2344
qed
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  2345
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2346
lemma compact_real_lemma:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2347
  assumes "\<forall>n::nat. abs(s n) \<le> b"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2348
  shows "\<exists>(l::real) r. subseq r \<and> ((s \<circ> r) ---> l) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2349
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2350
  obtain r where r:"subseq r" "monoseq (\<lambda>n. s (r n))"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2351
    using seq_monosub[of s] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2352
  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
  2353
    unfolding tendsto_iff dist_norm eventually_sequentially by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2354
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2355
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2356
instance real :: heine_borel
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2357
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2358
  fix s :: "real set" and f :: "nat \<Rightarrow> real"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2359
  assume s: "bounded s" and f: "\<forall>n. f n \<in> s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2360
  then obtain b where b: "\<forall>n. abs (f n) \<le> b"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2361
    unfolding bounded_iff by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2362
  obtain l :: real and r :: "nat \<Rightarrow> nat" where
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2363
    r: "subseq r" and l: "((f \<circ> r) ---> l) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2364
    using compact_real_lemma [OF b] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2365
  thus "\<exists>l r. subseq r \<and> ((f \<circ> r) ---> l) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2366
    by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2367
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2368
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2369
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
  2370
  fixes f :: "nat \<Rightarrow> 'a::euclidean_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2371
  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
  2372
  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
  2373
        (\<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
  2374
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
  2375
  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
  2376
  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
  2377
  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
  2378
      (\<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
  2379
  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
  2380
  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
  2381
    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
  2382
      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
  2383
    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
  2384
      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
  2385
      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
  2386
    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
  2387
    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
  2388
      using bounded_imp_convergent_subsequence[OF s' f'] unfolding o_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2389
    def r \<equiv> "r1 \<circ> r2" have r:"subseq r"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2390
      using r1 and r2 unfolding r_def o_def subseq_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2391
    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
  2392
    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
  2393
    { 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
  2394
      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
  2395
      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
  2396
      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
  2397
        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
  2398
      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
  2399
        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
  2400
        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
  2401
    }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2402
    ultimately show ?case by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2403
  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
  2404
qed
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  2405
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  2406
instance euclidean_space \<subseteq> heine_borel
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2407
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
  2408
  fix s :: "'a set" and f :: "nat \<Rightarrow> 'a"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2409
  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
  2410
  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
  2411
    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
  2412
    using compact_lemma [OF s f] by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2413
  { 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
  2414
    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
  2415
    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
  2416
      by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2417
    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
  2418
    { 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
  2419
      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
  2420
        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
  2421
      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
  2422
        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
  2423
      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
  2424
        by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2425
    }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2426
    ultimately have "eventually (\<lambda>n. dist (f (r n)) l < e) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2427
      by (rule eventually_elim1)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2428
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2429
  hence *:"((f \<circ> r) ---> l) sequentially" unfolding o_def tendsto_iff by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2430
  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
  2431
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2432
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2433
lemma bounded_fst: "bounded s \<Longrightarrow> bounded (fst ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2434
unfolding bounded_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2435
apply clarify
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2436
apply (rule_tac x="a" in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2437
apply (rule_tac x="e" in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2438
apply clarsimp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2439
apply (drule (1) bspec)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2440
apply (simp add: dist_Pair_Pair)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2441
apply (erule order_trans [OF real_sqrt_sum_squares_ge1])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2442
done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2443
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2444
lemma bounded_snd: "bounded s \<Longrightarrow> bounded (snd ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2445
unfolding bounded_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2446
apply clarify
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2447
apply (rule_tac x="b" in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2448
apply (rule_tac x="e" in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2449
apply clarsimp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2450
apply (drule (1) bspec)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2451
apply (simp add: dist_Pair_Pair)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2452
apply (erule order_trans [OF real_sqrt_sum_squares_ge2])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2453
done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2454
37678
0040bafffdef "prod" and "sum" replace "*" and "+" respectively
haftmann
parents: 37649
diff changeset
  2455
instance prod :: (heine_borel, heine_borel) heine_borel
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2456
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2457
  fix s :: "('a * 'b) set" and f :: "nat \<Rightarrow> 'a * 'b"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2458
  assume s: "bounded s" and f: "\<forall>n. f n \<in> s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2459
  from s have s1: "bounded (fst ` s)" by (rule bounded_fst)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2460
  from f have f1: "\<forall>n. fst (f n) \<in> fst ` s" by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2461
  obtain l1 r1 where r1: "subseq r1"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2462
    and l1: "((\<lambda>n. fst (f (r1 n))) ---> l1) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2463
    using bounded_imp_convergent_subsequence [OF s1 f1]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2464
    unfolding o_def by fast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2465
  from s have s2: "bounded (snd ` s)" by (rule bounded_snd)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2466
  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
  2467
  obtain l2 r2 where r2: "subseq r2"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2468
    and l2: "((\<lambda>n. snd (f (r1 (r2 n)))) ---> l2) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2469
    using bounded_imp_convergent_subsequence [OF s2 f2]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2470
    unfolding o_def by fast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2471
  have l1': "((\<lambda>n. fst (f (r1 (r2 n)))) ---> l1) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2472
    using lim_subseq [OF r2 l1] unfolding o_def .
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2473
  have l: "((f \<circ> (r1 \<circ> r2)) ---> (l1, l2)) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2474
    using tendsto_Pair [OF l1' l2] unfolding o_def by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2475
  have r: "subseq (r1 \<circ> r2)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2476
    using r1 r2 unfolding subseq_def by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2477
  show "\<exists>l r. subseq r \<and> ((f \<circ> r) ---> l) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2478
    using l r by fast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2479
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2480
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  2481
subsubsection{* Completeness *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2482
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2483
lemma cauchy_def:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2484
  "Cauchy s \<longleftrightarrow> (\<forall>e>0. \<exists>N. \<forall>m n. m \<ge> N \<and> n \<ge> N --> dist(s m)(s n) < e)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2485
unfolding Cauchy_def by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2486
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2487
definition
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2488
  complete :: "'a::metric_space set \<Rightarrow> bool" where
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2489
  "complete s \<longleftrightarrow> (\<forall>f. (\<forall>n. f n \<in> s) \<and> Cauchy f
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2490
                      --> (\<exists>l \<in> s. (f ---> l) sequentially))"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2491
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2492
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
  2493
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2494
  { assume ?rhs
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2495
    { fix e::real
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2496
      assume "e>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2497
      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
  2498
        by (erule_tac x="e/2" in allE) auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2499
      { fix n m
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2500
        assume nm:"N \<le> m \<and> N \<le> n"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2501
        hence "dist (s m) (s n) < e" using N
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2502
          using dist_triangle_half_l[of "s m" "s N" "e" "s n"]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2503
          by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2504
      }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2505
      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
  2506
        by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2507
    }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2508
    hence ?lhs
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2509
      unfolding cauchy_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2510
      by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2511
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2512
  thus ?thesis
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2513
    unfolding cauchy_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2514
    using dist_triangle_half_l
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2515
    by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2516
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2517
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2518
lemma convergent_imp_cauchy:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2519
 "(s ---> l) sequentially ==> Cauchy s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2520
proof(simp only: cauchy_def, rule, rule)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2521
  fix e::real assume "e>0" "(s ---> l) sequentially"
44907
93943da0a010 remove redundant lemma Lim_sequentially in favor of lemma LIMSEQ_def
huffman
parents: 44905
diff changeset
  2522
  then obtain N::nat where N:"\<forall>n\<ge>N. dist (s n) l < e/2" unfolding LIMSEQ_def by(erule_tac x="e/2" in allE) auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2523
  thus "\<exists>N. \<forall>m n. N \<le> m \<and> N \<le> n \<longrightarrow> dist (s m) (s n) < e"  using dist_triangle_half_l[of _ l e _] by (rule_tac x=N in exI) auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2524
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2525
34104
22758f95e624 re-state lemmas using 'range'
huffman
parents: 33758
diff changeset
  2526
lemma cauchy_imp_bounded: assumes "Cauchy s" shows "bounded (range s)"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2527
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2528
  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
  2529
  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
  2530
  moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2531
  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
  2532
  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
  2533
    unfolding bounded_any_center [where a="s N"] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2534
  ultimately show "?thesis"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2535
    unfolding bounded_any_center [where a="s N"]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2536
    apply(rule_tac x="max a 1" in exI) apply auto
34104
22758f95e624 re-state lemmas using 'range'
huffman
parents: 33758
diff changeset
  2537
    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
  2538
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2539
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2540
lemma compact_imp_complete: assumes "compact s" shows "complete s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2541
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2542
  { fix f assume as: "(\<forall>n::nat. f n \<in> s)" "Cauchy f"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2543
    from as(1) obtain l r where lr: "l\<in>s" "subseq r" "((f \<circ> r) ---> l) sequentially" using assms unfolding compact_def by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2544
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2545
    note lr' = subseq_bigger [OF lr(2)]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2546
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2547
    { fix e::real assume "e>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2548
      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
  2549
      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
  2550
      { fix n::nat assume n:"n \<ge> max N M"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2551
        have "dist ((f \<circ> r) n) l < e/2" using n M by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2552
        moreover have "r n \<ge> N" using lr'[of n] n by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2553
        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
  2554
        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
  2555
      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
  2556
    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
  2557
  thus ?thesis unfolding complete_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2558
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2559
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2560
instance heine_borel < complete_space
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2561
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2562
  fix f :: "nat \<Rightarrow> 'a" assume "Cauchy f"
34104
22758f95e624 re-state lemmas using 'range'
huffman
parents: 33758
diff changeset
  2563
  hence "bounded (range f)"
22758f95e624 re-state lemmas using 'range'
huffman
parents: 33758
diff changeset
  2564
    by (rule cauchy_imp_bounded)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2565
  hence "compact (closure (range f))"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2566
    using bounded_closed_imp_compact [of "closure (range f)"] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2567
  hence "complete (closure (range f))"
34104
22758f95e624 re-state lemmas using 'range'
huffman
parents: 33758
diff changeset
  2568
    by (rule compact_imp_complete)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2569
  moreover have "\<forall>n. f n \<in> closure (range f)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2570
    using closure_subset [of "range f"] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2571
  ultimately have "\<exists>l\<in>closure (range f). (f ---> l) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2572
    using `Cauchy f` unfolding complete_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2573
  then show "convergent f"
36660
1cc4ab4b7ff7 make (X ----> L) an abbreviation for (X ---> L) sequentially
huffman
parents: 36659
diff changeset
  2574
    unfolding convergent_def by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2575
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2576
44632
076a45f65e12 simplify/generalize some proofs
huffman
parents: 44628
diff changeset
  2577
instance euclidean_space \<subseteq> banach ..
076a45f65e12 simplify/generalize some proofs
huffman
parents: 44628
diff changeset
  2578
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2579
lemma complete_univ: "complete (UNIV :: 'a::complete_space set)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2580
proof(simp add: complete_def, rule, rule)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2581
  fix f :: "nat \<Rightarrow> 'a" assume "Cauchy f"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2582
  hence "convergent f" by (rule Cauchy_convergent)
36660
1cc4ab4b7ff7 make (X ----> L) an abbreviation for (X ---> L) sequentially
huffman
parents: 36659
diff changeset
  2583
  thus "\<exists>l. f ----> l" unfolding convergent_def .  
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2584
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2585
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2586
lemma complete_imp_closed: assumes "complete s" shows "closed s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2587
proof -
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2588
  { fix x assume "x islimpt s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2589
    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
  2590
      unfolding islimpt_sequential by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2591
    then obtain l where l: "l\<in>s" "(f ---> l) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2592
      using `complete s`[unfolded complete_def] using convergent_imp_cauchy[of f x] by auto
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41969
diff changeset
  2593
    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
  2594
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2595
  thus "closed s" unfolding closed_limpt by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2596
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2597
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2598
lemma complete_eq_closed:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2599
  fixes s :: "'a::complete_space set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2600
  shows "complete s \<longleftrightarrow> closed s" (is "?lhs = ?rhs")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2601
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2602
  assume ?lhs thus ?rhs by (rule complete_imp_closed)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2603
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2604
  assume ?rhs
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2605
  { fix f assume as:"\<forall>n::nat. f n \<in> s" "Cauchy f"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2606
    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
  2607
    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
  2608
  thus ?lhs unfolding complete_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2609
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2610
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2611
lemma convergent_eq_cauchy:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2612
  fixes s :: "nat \<Rightarrow> 'a::complete_space"
44632
076a45f65e12 simplify/generalize some proofs
huffman
parents: 44628
diff changeset
  2613
  shows "(\<exists>l. (s ---> l) sequentially) \<longleftrightarrow> Cauchy s"
076a45f65e12 simplify/generalize some proofs
huffman
parents: 44628
diff changeset
  2614
  unfolding Cauchy_convergent_iff convergent_def ..
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2615
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2616
lemma convergent_imp_bounded:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2617
  fixes s :: "nat \<Rightarrow> 'a::metric_space"
44632
076a45f65e12 simplify/generalize some proofs
huffman
parents: 44628
diff changeset
  2618
  shows "(s ---> l) sequentially \<Longrightarrow> bounded (range s)"
076a45f65e12 simplify/generalize some proofs
huffman
parents: 44628
diff changeset
  2619
  by (intro cauchy_imp_bounded convergent_imp_cauchy)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2620
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  2621
subsubsection{* Total boundedness *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2622
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2623
fun helper_1::"('a::metric_space set) \<Rightarrow> real \<Rightarrow> nat \<Rightarrow> 'a" where
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2624
  "helper_1 s e n = (SOME y::'a. y \<in> s \<and> (\<forall>m<n. \<not> (dist (helper_1 s e m) y < e)))"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2625
declare helper_1.simps[simp del]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2626
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2627
lemma compact_imp_totally_bounded:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2628
  assumes "compact s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2629
  shows "\<forall>e>0. \<exists>k. finite k \<and> k \<subseteq> s \<and> s \<subseteq> (\<Union>((\<lambda>x. ball x e) ` k))"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2630
proof(rule, rule, rule ccontr)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2631
  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)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2632
  def x \<equiv> "helper_1 s e"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2633
  { fix n
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2634
    have "x n \<in> s \<and> (\<forall>m<n. \<not> dist (x m) (x n) < e)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2635
    proof(induct_tac rule:nat_less_induct)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2636
      fix n  def Q \<equiv> "(\<lambda>y. y \<in> s \<and> (\<forall>m<n. \<not> dist (x m) y < e))"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2637
      assume as:"\<forall>m<n. x m \<in> s \<and> (\<forall>ma<m. \<not> dist (x ma) (x m) < e)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2638
      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
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2639
      then obtain z where z:"z\<in>s" "z \<notin> (\<Union>x\<in>x ` {0..<n}. ball x e)" unfolding subset_eq by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2640
      have "Q (x n)" unfolding x_def and helper_1.simps[of s e n]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2641
        apply(rule someI2[where a=z]) unfolding x_def[symmetric] and Q_def using z by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2642
      thus "x n \<in> s \<and> (\<forall>m<n. \<not> dist (x m) (x n) < e)" unfolding Q_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2643
    qed }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2644
  hence "\<forall>n::nat. x n \<in> s" and x:"\<forall>n. \<forall>m < n. \<not> (dist (x m) (x n) < e)" by blast+
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2645
  then obtain l r where "l\<in>s" and r:"subseq r" and "((x \<circ> r) ---> l) sequentially" using assms(1)[unfolded compact_def, THEN spec[where x=x]] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2646
  from this(3) have "Cauchy (x \<circ> r)" using convergent_imp_cauchy by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2647
  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
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2648
  show False
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2649
    using N[THEN spec[where x=N], THEN spec[where x="N+1"]]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2650
    using r[unfolded subseq_def, THEN spec[where x=N], THEN spec[where x="N+1"]]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2651
    using x[THEN spec[where x="r (N+1)"], THEN spec[where x="r (N)"]] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2652
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2653
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  2654
subsubsection{* Heine-Borel theorem *}
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  2655
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  2656
text {* Following Burkill \& Burkill vol. 2. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2657
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2658
lemma heine_borel_lemma: fixes s::"'a::metric_space set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2659
  assumes "compact s"  "s \<subseteq> (\<Union> t)"  "\<forall>b \<in> t. open b"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2660
  shows "\<exists>e>0. \<forall>x \<in> s. \<exists>b \<in> t. ball x e \<subseteq> b"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2661
proof(rule ccontr)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2662
  assume "\<not> (\<exists>e>0. \<forall>x\<in>s. \<exists>b\<in>t. ball x e \<subseteq> b)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2663
  hence cont:"\<forall>e>0. \<exists>x\<in>s. \<forall>xa\<in>t. \<not> (ball x e \<subseteq> xa)" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2664
  { fix n::nat
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2665
    have "1 / real (n + 1) > 0" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2666
    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 }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2667
  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
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2668
  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)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2669
    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
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2670
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2671
  then obtain l r where l:"l\<in>s" and r:"subseq r" and lr:"((f \<circ> r) ---> l) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2672
    using assms(1)[unfolded compact_def, THEN spec[where x=f]] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2673
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2674
  obtain b where "l\<in>b" "b\<in>t" using assms(2) and l by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2675
  then obtain e where "e>0" and e:"\<forall>z. dist z l < e \<longrightarrow> z\<in>b"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2676
    using assms(3)[THEN bspec[where x=b]] unfolding open_dist by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2677
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2678
  then obtain N1 where N1:"\<forall>n\<ge>N1. dist ((f \<circ> r) n) l < e / 2"
44907
93943da0a010 remove redundant lemma Lim_sequentially in favor of lemma LIMSEQ_def
huffman
parents: 44905
diff changeset
  2679
    using lr[unfolded LIMSEQ_def, THEN spec[where x="e/2"]] by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2680
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2681
  obtain N2::nat where N2:"N2>0" "inverse (real N2) < e /2" using real_arch_inv[of "e/2"] and `e>0` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2682
  have N2':"inverse (real (r (N1 + N2) +1 )) < e/2"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2683
    apply(rule order_less_trans) apply(rule less_imp_inverse_less) using N2
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2684
    using subseq_bigger[OF r, of "N1 + N2"] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2685
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2686
  def x \<equiv> "(f (r (N1 + N2)))"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2687
  have x:"\<not> ball x (inverse (real (r (N1 + N2) + 1))) \<subseteq> b" unfolding x_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2688
    using f[THEN spec[where x="r (N1 + N2)"]] using `b\<in>t` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2689
  have "\<exists>y\<in>ball x (inverse (real (r (N1 + N2) + 1))). y\<notin>b" apply(rule ccontr) using x by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2690
  then obtain y where y:"y \<in> ball x (inverse (real (r (N1 + N2) + 1)))" "y \<notin> b" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2691
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2692
  have "dist x l < e/2" using N1 unfolding x_def o_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2693
  hence "dist y l < e" using y N2' using dist_triangle[of y l x]by (auto simp add:dist_commute)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2694
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2695
  thus False using e and `y\<notin>b` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2696
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2697
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2698
lemma compact_imp_heine_borel: "compact s ==> (\<forall>f. (\<forall>t \<in> f. open t) \<and> s \<subseteq> (\<Union> f)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2699
               \<longrightarrow> (\<exists>f'. f' \<subseteq> f \<and> finite f' \<and> s \<subseteq> (\<Union> f')))"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2700
proof clarify
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2701
  fix f assume "compact s" " \<forall>t\<in>f. open t" "s \<subseteq> \<Union>f"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2702
  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
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2703
  hence "\<forall>x\<in>s. \<exists>b. b\<in>f \<and> ball x e \<subseteq> b" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2704
  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
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2705
  then obtain  bb where bb:"\<forall>x\<in>s. (bb x) \<in> f \<and> ball x e \<subseteq> (bb x)" by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2706
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2707
  from `compact s` have  "\<exists> k. finite k \<and> k \<subseteq> s \<and> s \<subseteq> \<Union>(\<lambda>x. ball x e) ` k" using compact_imp_totally_bounded[of s] `e>0` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2708
  then obtain k where k:"finite k" "k \<subseteq> s" "s \<subseteq> \<Union>(\<lambda>x. ball x e) ` k" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2709
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2710
  have "finite (bb ` k)" using k(1) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2711
  moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2712
  { fix x assume "x\<in>s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2713
    hence "x\<in>\<Union>(\<lambda>x. ball x e) ` k" using k(3)  unfolding subset_eq by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2714
    hence "\<exists>X\<in>bb ` k. x \<in> X" using bb k(2) by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2715
    hence "x \<in> \<Union>(bb ` k)" using  Union_iff[of x "bb ` k"] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2716
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2717
  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
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2718
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2719
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  2720
subsubsection {* Bolzano-Weierstrass property *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2721
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2722
lemma heine_borel_imp_bolzano_weierstrass:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2723
  assumes "\<forall>f. (\<forall>t \<in> f. open t) \<and> s \<subseteq> (\<Union> f) --> (\<exists>f'. f' \<subseteq> f \<and> finite f' \<and> s \<subseteq> (\<Union> f'))"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2724
          "infinite t"  "t \<subseteq> s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2725
  shows "\<exists>x \<in> s. x islimpt t"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2726
proof(rule ccontr)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2727
  assume "\<not> (\<exists>x \<in> s. x islimpt t)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2728
  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
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2729
    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
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2730
  obtain g where g:"g\<subseteq>{t. \<exists>x. x \<in> s \<and> t = f x}" "finite g" "s \<subseteq> \<Union>g"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2731
    using assms(1)[THEN spec[where x="{t. \<exists>x. x\<in>s \<and> t = f x}"]] using f by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2732
  from g(1,3) have g':"\<forall>x\<in>g. \<exists>xa \<in> s. x = f xa" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2733
  { fix x y assume "x\<in>t" "y\<in>t" "f x = f y"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2734
    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
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2735
    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  }
36362
06475a1547cb fix lots of looping simp calls and other warnings
huffman
parents: 36360
diff changeset
  2736
  hence "inj_on f t" unfolding inj_on_def by simp
06475a1547cb fix lots of looping simp calls and other warnings
huffman
parents: 36360
diff changeset
  2737
  hence "infinite (f ` t)" using assms(2) using finite_imageD by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2738
  moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2739
  { fix x assume "x\<in>t" "f x \<notin> g"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2740
    from g(3) assms(3) `x\<in>t` obtain h where "h\<in>g" and "x\<in>h" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2741
    then obtain y where "y\<in>s" "h = f y" using g'[THEN bspec[where x=h]] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2742
    hence "y = x" using f[THEN bspec[where x=y]] and `x\<in>t` and `x\<in>h`[unfolded `h = f y`] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2743
    hence False using `f x \<notin> g` `h\<in>g` unfolding `h = f y` by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2744
  hence "f ` t \<subseteq> g" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2745
  ultimately show False using g(2) using finite_subset by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2746
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2747
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  2748
subsubsection {* Complete the chain of compactness variants *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2749
44073
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2750
lemma islimpt_range_imp_convergent_subsequence:
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2751
  fixes f :: "nat \<Rightarrow> 'a::metric_space"
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2752
  assumes "l islimpt (range f)"
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2753
  shows "\<exists>r. subseq r \<and> ((f \<circ> r) ---> l) sequentially"
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2754
proof (intro exI conjI)
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2755
  have *: "\<And>e. 0 < e \<Longrightarrow> \<exists>n. 0 < dist (f n) l \<and> dist (f n) l < e"
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2756
    using assms unfolding islimpt_def
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2757
    by (drule_tac x="ball l e" in spec)
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2758
       (auto simp add: zero_less_dist_iff dist_commute)
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2759
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2760
  def t \<equiv> "\<lambda>e. LEAST n. 0 < dist (f n) l \<and> dist (f n) l < e"
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2761
  have f_t_neq: "\<And>e. 0 < e \<Longrightarrow> 0 < dist (f (t e)) l"
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2762
    unfolding t_def by (rule LeastI2_ex [OF * conjunct1])
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2763
  have f_t_closer: "\<And>e. 0 < e \<Longrightarrow> dist (f (t e)) l < e"
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2764
    unfolding t_def by (rule LeastI2_ex [OF * conjunct2])
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2765
  have t_le: "\<And>n e. 0 < dist (f n) l \<Longrightarrow> dist (f n) l < e \<Longrightarrow> t e \<le> n"
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2766
    unfolding t_def by (simp add: Least_le)
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2767
  have less_tD: "\<And>n e. n < t e \<Longrightarrow> 0 < dist (f n) l \<Longrightarrow> e \<le> dist (f n) l"
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2768
    unfolding t_def by (drule not_less_Least) simp
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2769
  have t_antimono: "\<And>e e'. 0 < e \<Longrightarrow> e \<le> e' \<Longrightarrow> t e' \<le> t e"
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2770
    apply (rule t_le)
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2771
    apply (erule f_t_neq)
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2772
    apply (erule (1) less_le_trans [OF f_t_closer])
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2773
    done
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2774
  have t_dist_f_neq: "\<And>n. 0 < dist (f n) l \<Longrightarrow> t (dist (f n) l) \<noteq> n"
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2775
    by (drule f_t_closer) auto
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2776
  have t_less: "\<And>e. 0 < e \<Longrightarrow> t e < t (dist (f (t e)) l)"
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2777
    apply (subst less_le)
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2778
    apply (rule conjI)
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2779
    apply (rule t_antimono)
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2780
    apply (erule f_t_neq)
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2781
    apply (erule f_t_closer [THEN less_imp_le])
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2782
    apply (rule t_dist_f_neq [symmetric])
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2783
    apply (erule f_t_neq)
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2784
    done
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2785
  have dist_f_t_less':
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2786
    "\<And>e e'. 0 < e \<Longrightarrow> 0 < e' \<Longrightarrow> t e \<le> t e' \<Longrightarrow> dist (f (t e')) l < e"
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2787
    apply (simp add: le_less)
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2788
    apply (erule disjE)
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2789
    apply (rule less_trans)
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2790
    apply (erule f_t_closer)
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2791
    apply (rule le_less_trans)
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2792
    apply (erule less_tD)
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2793
    apply (erule f_t_neq)
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2794
    apply (erule f_t_closer)
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2795
    apply (erule subst)
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2796
    apply (erule f_t_closer)
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2797
    done
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2798
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2799
  def r \<equiv> "nat_rec (t 1) (\<lambda>_ x. t (dist (f x) l))"
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2800
  have r_simps: "r 0 = t 1" "\<And>n. r (Suc n) = t (dist (f (r n)) l)"
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2801
    unfolding r_def by simp_all
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2802
  have f_r_neq: "\<And>n. 0 < dist (f (r n)) l"
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2803
    by (induct_tac n) (simp_all add: r_simps f_t_neq)
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2804
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2805
  show "subseq r"
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2806
    unfolding subseq_Suc_iff
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2807
    apply (rule allI)
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2808
    apply (case_tac n)
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2809
    apply (simp_all add: r_simps)
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2810
    apply (rule t_less, rule zero_less_one)
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2811
    apply (rule t_less, rule f_r_neq)
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2812
    done
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2813
  show "((f \<circ> r) ---> l) sequentially"
44907
93943da0a010 remove redundant lemma Lim_sequentially in favor of lemma LIMSEQ_def
huffman
parents: 44905
diff changeset
  2814
    unfolding LIMSEQ_def o_def
93943da0a010 remove redundant lemma Lim_sequentially in favor of lemma LIMSEQ_def
huffman
parents: 44905
diff changeset
  2815
    apply (clarify, rename_tac e, rule_tac x="t e" in exI, clarify)
44073
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2816
    apply (drule le_trans, rule seq_suble [OF `subseq r`])
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2817
    apply (case_tac n, simp_all add: r_simps dist_f_t_less' f_r_neq)
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2818
    done
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2819
qed
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2820
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2821
lemma finite_range_imp_infinite_repeats:
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2822
  fixes f :: "nat \<Rightarrow> 'a"
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2823
  assumes "finite (range f)"
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2824
  shows "\<exists>k. infinite {n. f n = k}"
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2825
proof -
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2826
  { fix A :: "'a set" assume "finite A"
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2827
    hence "\<And>f. infinite {n. f n \<in> A} \<Longrightarrow> \<exists>k. infinite {n. f n = k}"
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2828
    proof (induct)
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2829
      case empty thus ?case by simp
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2830
    next
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2831
      case (insert x A)
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2832
     show ?case
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2833
      proof (cases "finite {n. f n = x}")
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2834
        case True
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2835
        with `infinite {n. f n \<in> insert x A}`
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2836
        have "infinite {n. f n \<in> A}" by simp
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2837
        thus "\<exists>k. infinite {n. f n = k}" by (rule insert)
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2838
      next
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2839
        case False thus "\<exists>k. infinite {n. f n = k}" ..
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2840
      qed
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2841
    qed
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2842
  } note H = this
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2843
  from assms show "\<exists>k. infinite {n. f n = k}"
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2844
    by (rule H) simp
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2845
qed
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2846
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2847
lemma bolzano_weierstrass_imp_compact:
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2848
  fixes s :: "'a::metric_space set"
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2849
  assumes "\<forall>t. infinite t \<and> t \<subseteq> s --> (\<exists>x \<in> s. x islimpt t)"
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2850
  shows "compact s"
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2851
proof -
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2852
  { fix f :: "nat \<Rightarrow> 'a" assume f: "\<forall>n. f n \<in> s"
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2853
    have "\<exists>l\<in>s. \<exists>r. subseq r \<and> ((f \<circ> r) ---> l) sequentially"
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2854
    proof (cases "finite (range f)")
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2855
      case True
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2856
      hence "\<exists>l. infinite {n. f n = l}"
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2857
        by (rule finite_range_imp_infinite_repeats)
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2858
      then obtain l where "infinite {n. f n = l}" ..
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2859
      hence "\<exists>r. subseq r \<and> (\<forall>n. r n \<in> {n. f n = l})"
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2860
        by (rule infinite_enumerate)
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2861
      then obtain r where "subseq r" and fr: "\<forall>n. f (r n) = l" by auto
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2862
      hence "subseq r \<and> ((f \<circ> r) ---> l) sequentially"
44125
230a8665c919 mark some redundant theorems as legacy
huffman
parents: 44122
diff changeset
  2863
        unfolding o_def by (simp add: fr tendsto_const)
44073
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2864
      hence "\<exists>r. subseq r \<and> ((f \<circ> r) ---> l) sequentially"
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2865
        by - (rule exI)
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2866
      from f have "\<forall>n. f (r n) \<in> s" by simp
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2867
      hence "l \<in> s" by (simp add: fr)
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2868
      thus "\<exists>l\<in>s. \<exists>r. subseq r \<and> ((f \<circ> r) ---> l) sequentially"
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2869
        by (rule rev_bexI) fact
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2870
    next
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2871
      case False
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2872
      with f assms have "\<exists>x\<in>s. x islimpt (range f)" by auto
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2873
      then obtain l where "l \<in> s" "l islimpt (range f)" ..
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2874
      have "\<exists>r. subseq r \<and> ((f \<circ> r) ---> l) sequentially"
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2875
        using `l islimpt (range f)`
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2876
        by (rule islimpt_range_imp_convergent_subsequence)
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2877
      with `l \<in> s` show "\<exists>l\<in>s. \<exists>r. subseq r \<and> ((f \<circ> r) ---> l) sequentially" ..
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2878
    qed
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2879
  }
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2880
  thus ?thesis unfolding compact_def by auto
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2881
qed
efd1ea744101 lemma bolzano_weierstrass_imp_compact
huffman
parents: 44072
diff changeset
  2882
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2883
primrec helper_2::"(real \<Rightarrow> 'a::metric_space) \<Rightarrow> nat \<Rightarrow> 'a" where
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2884
  "helper_2 beyond 0 = beyond 0" |
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2885
  "helper_2 beyond (Suc n) = beyond (dist undefined (helper_2 beyond n) + 1 )"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2886
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2887
lemma bolzano_weierstrass_imp_bounded: fixes s::"'a::metric_space set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2888
  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
  2889
  shows "bounded s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2890
proof(rule ccontr)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2891
  assume "\<not> bounded s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2892
  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
  2893
    unfolding bounded_any_center [where a=undefined]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2894
    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
  2895
  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
  2896
    unfolding linorder_not_le by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2897
  def x \<equiv> "helper_2 beyond"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2898
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2899
  { fix m n ::nat assume "m<n"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2900
    hence "dist undefined (x m) + 1 < dist undefined (x n)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2901
    proof(induct n)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2902
      case 0 thus ?case by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2903
    next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2904
      case (Suc n)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2905
      have *:"dist undefined (x n) + 1 < dist undefined (x (Suc n))"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2906
        unfolding x_def and helper_2.simps
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2907
        using beyond(2)[of "dist undefined (helper_2 beyond n) + 1"] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2908
      thus ?case proof(cases "m < n")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2909
        case True thus ?thesis using Suc and * by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2910
      next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2911
        case False hence "m = n" using Suc(2) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2912
        thus ?thesis using * by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2913
      qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2914
    qed  } note * = this
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2915
  { fix m n ::nat assume "m\<noteq>n"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2916
    have "1 < dist (x m) (x n)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2917
    proof(cases "m<n")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2918
      case True
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2919
      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
  2920
      thus ?thesis using dist_triangle [of undefined "x n" "x m"] by arith
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2921
    next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2922
      case False hence "n<m" using `m\<noteq>n` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2923
      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
  2924
      thus ?thesis using dist_triangle2 [of undefined "x m" "x n"] by arith
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2925
    qed  } note ** = this
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2926
  { fix a b assume "x a = x b" "a \<noteq> b"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2927
    hence False using **[of a b] by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2928
  hence "inj x" unfolding inj_on_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2929
  moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2930
  { fix n::nat
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2931
    have "x n \<in> s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2932
    proof(cases "n = 0")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2933
      case True thus ?thesis unfolding x_def using beyond by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2934
    next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2935
      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
  2936
      thus ?thesis unfolding x_def using beyond by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2937
    qed  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2938
  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
  2939
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2940
  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
  2941
  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
  2942
  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
  2943
    unfolding dist_nz by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2944
  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
  2945
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2946
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2947
lemma sequence_infinite_lemma:
44076
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  2948
  fixes f :: "nat \<Rightarrow> 'a::t1_space"
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  2949
  assumes "\<forall>n. f n \<noteq> l" and "(f ---> l) sequentially"
34104
22758f95e624 re-state lemmas using 'range'
huffman
parents: 33758
diff changeset
  2950
  shows "infinite (range f)"
22758f95e624 re-state lemmas using 'range'
huffman
parents: 33758
diff changeset
  2951
proof
22758f95e624 re-state lemmas using 'range'
huffman
parents: 33758
diff changeset
  2952
  assume "finite (range f)"
44076
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  2953
  hence "closed (range f)" by (rule finite_imp_closed)
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  2954
  hence "open (- range f)" by (rule open_Compl)
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  2955
  from assms(1) have "l \<in> - range f" by auto
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  2956
  from assms(2) have "eventually (\<lambda>n. f n \<in> - range f) sequentially"
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  2957
    using `open (- range f)` `l \<in> - range f` by (rule topological_tendstoD)
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  2958
  thus False unfolding eventually_sequentially by auto
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  2959
qed
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  2960
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  2961
lemma closure_insert:
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  2962
  fixes x :: "'a::t1_space"
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  2963
  shows "closure (insert x s) = insert x (closure s)"
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  2964
apply (rule closure_unique)
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  2965
apply (rule insert_mono [OF closure_subset])
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  2966
apply (rule closed_insert [OF closed_closure])
44076
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  2967
apply (simp add: closure_minimal)
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  2968
done
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  2969
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  2970
lemma islimpt_insert:
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  2971
  fixes x :: "'a::t1_space"
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  2972
  shows "x islimpt (insert a s) \<longleftrightarrow> x islimpt s"
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  2973
proof
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  2974
  assume *: "x islimpt (insert a s)"
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  2975
  show "x islimpt s"
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  2976
  proof (rule islimptI)
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  2977
    fix t assume t: "x \<in> t" "open t"
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  2978
    show "\<exists>y\<in>s. y \<in> t \<and> y \<noteq> x"
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  2979
    proof (cases "x = a")
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  2980
      case True
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  2981
      obtain y where "y \<in> insert a s" "y \<in> t" "y \<noteq> x"
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  2982
        using * t by (rule islimptE)
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  2983
      with `x = a` show ?thesis by auto
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  2984
    next
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  2985
      case False
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  2986
      with t have t': "x \<in> t - {a}" "open (t - {a})"
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  2987
        by (simp_all add: open_Diff)
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  2988
      obtain y where "y \<in> insert a s" "y \<in> t - {a}" "y \<noteq> x"
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  2989
        using * t' by (rule islimptE)
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  2990
      thus ?thesis by auto
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  2991
    qed
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  2992
  qed
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  2993
next
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  2994
  assume "x islimpt s" thus "x islimpt (insert a s)"
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  2995
    by (rule islimpt_subset) auto
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  2996
qed
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  2997
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  2998
lemma islimpt_union_finite:
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  2999
  fixes x :: "'a::t1_space"
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  3000
  shows "finite s \<Longrightarrow> x islimpt (s \<union> t) \<longleftrightarrow> x islimpt t"
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  3001
by (induct set: finite, simp_all add: islimpt_insert)
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  3002
 
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3003
lemma sequence_unique_limpt:
44076
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  3004
  fixes f :: "nat \<Rightarrow> 'a::t2_space"
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  3005
  assumes "(f ---> l) sequentially" and "l' islimpt (range f)"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3006
  shows "l' = l"
44076
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  3007
proof (rule ccontr)
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  3008
  assume "l' \<noteq> l"
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  3009
  obtain s t where "open s" "open t" "l' \<in> s" "l \<in> t" "s \<inter> t = {}"
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  3010
    using hausdorff [OF `l' \<noteq> l`] by auto
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  3011
  have "eventually (\<lambda>n. f n \<in> t) sequentially"
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  3012
    using assms(1) `open t` `l \<in> t` by (rule topological_tendstoD)
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  3013
  then obtain N where "\<forall>n\<ge>N. f n \<in> t"
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  3014
    unfolding eventually_sequentially by auto
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  3015
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  3016
  have "UNIV = {..<N} \<union> {N..}" by auto
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  3017
  hence "l' islimpt (f ` ({..<N} \<union> {N..}))" using assms(2) by simp
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  3018
  hence "l' islimpt (f ` {..<N} \<union> f ` {N..})" by (simp add: image_Un)
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  3019
  hence "l' islimpt (f ` {N..})" by (simp add: islimpt_union_finite)
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  3020
  then obtain y where "y \<in> f ` {N..}" "y \<in> s" "y \<noteq> l'"
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  3021
    using `l' \<in> s` `open s` by (rule islimptE)
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  3022
  then obtain n where "N \<le> n" "f n \<in> s" "f n \<noteq> l'" by auto
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  3023
  with `\<forall>n\<ge>N. f n \<in> t` have "f n \<in> s \<inter> t" by simp
cddb05f94183 generalize sequence lemmas
huffman
parents: 44075
diff changeset
  3024
  with `s \<inter> t = {}` show False by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3025
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3026
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3027
lemma bolzano_weierstrass_imp_closed:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3028
  fixes s :: "'a::metric_space set" (* TODO: can this be generalized? *)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3029
  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
  3030
  shows "closed s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3031
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3032
  { fix x l assume as: "\<forall>n::nat. x n \<in> s" "(x ---> l) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3033
    hence "l \<in> s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3034
    proof(cases "\<forall>n. x n \<noteq> l")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3035
      case False thus "l\<in>s" using as(1) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3036
    next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3037
      case True note cas = this
34104
22758f95e624 re-state lemmas using 'range'
huffman
parents: 33758
diff changeset
  3038
      with as(2) have "infinite (range x)" using sequence_infinite_lemma[of x l] by auto
22758f95e624 re-state lemmas using 'range'
huffman
parents: 33758
diff changeset
  3039
      then obtain l' where "l'\<in>s" "l' islimpt (range x)" using assms[THEN spec[where x="range x"]] as(1) by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3040
      thus "l\<in>s" using sequence_unique_limpt[of x l l'] using as cas by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3041
    qed  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3042
  thus ?thesis unfolding closed_sequential_limits by fast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3043
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3044
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  3045
text {* Hence express everything as an equivalence. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3046
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3047
lemma compact_eq_heine_borel:
44074
e3a744a157d4 generalize compactness equivalence lemmas
huffman
parents: 44073
diff changeset
  3048
  fixes s :: "'a::metric_space set"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3049
  shows "compact s \<longleftrightarrow>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3050
           (\<forall>f. (\<forall>t \<in> f. open t) \<and> s \<subseteq> (\<Union> f)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3051
               --> (\<exists>f'. f' \<subseteq> f \<and> finite f' \<and> s \<subseteq> (\<Union> f')))" (is "?lhs = ?rhs")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3052
proof
44074
e3a744a157d4 generalize compactness equivalence lemmas
huffman
parents: 44073
diff changeset
  3053
  assume ?lhs thus ?rhs by (rule compact_imp_heine_borel)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3054
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3055
  assume ?rhs
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3056
  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
  3057
    by (blast intro: heine_borel_imp_bolzano_weierstrass[of s])
44074
e3a744a157d4 generalize compactness equivalence lemmas
huffman
parents: 44073
diff changeset
  3058
  thus ?lhs by (rule bolzano_weierstrass_imp_compact)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3059
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3060
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3061
lemma compact_eq_bolzano_weierstrass:
44074
e3a744a157d4 generalize compactness equivalence lemmas
huffman
parents: 44073
diff changeset
  3062
  fixes s :: "'a::metric_space set"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3063
  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
  3064
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3065
  assume ?lhs thus ?rhs unfolding compact_eq_heine_borel using heine_borel_imp_bolzano_weierstrass[of s] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3066
next
44074
e3a744a157d4 generalize compactness equivalence lemmas
huffman
parents: 44073
diff changeset
  3067
  assume ?rhs thus ?lhs by (rule bolzano_weierstrass_imp_compact)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3068
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3069
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3070
lemma nat_approx_posE:
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3071
  fixes e::real
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3072
  assumes "0 < e"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3073
  obtains n::nat where "1 / (Suc n) < e"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3074
proof atomize_elim
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3075
  have " 1 / real (Suc (nat (ceiling (1/e)))) < 1 / (ceiling (1/e))"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3076
    by (rule divide_strict_left_mono) (auto intro!: mult_pos_pos simp: `0 < e`)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3077
  also have "1 / (ceiling (1/e)) \<le> 1 / (1/e)"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3078
    by (rule divide_left_mono) (auto intro!: divide_pos_pos simp: `0 < e`)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3079
  also have "\<dots> = e" by simp
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3080
  finally show  "\<exists>n. 1 / real (Suc n) < e" ..
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3081
qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3082
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3083
lemma compact_eq_totally_bounded:
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3084
  shows "compact s \<longleftrightarrow> complete s \<and> (\<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
  3085
proof (safe intro!: compact_imp_complete)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3086
  fix e::real
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3087
  def f \<equiv> "(\<lambda>x::'a. ball x e) ` UNIV"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3088
  assume "0 < e" "compact s"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3089
  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
  3090
    by (simp add: compact_eq_heine_borel)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3091
  moreover
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3092
  have d0: "\<And>x::'a. dist x x < e" using `0 < e` by simp
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3093
  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
  3094
  ultimately have "(\<exists>f'\<subseteq>f. finite f' \<and> s \<subseteq> \<Union>f')" ..
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3095
  then guess K .. note K = this
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3096
  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
  3097
  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
  3098
  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
  3099
    by (intro exI[where x="k ` K"]) (auto simp: f_def)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3100
next
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3101
  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
  3102
  show "compact s"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3103
  proof cases
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3104
    assume "s = {}" thus "compact s" by (simp add: compact_def)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3105
  next
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3106
    assume "s \<noteq> {}"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3107
    show ?thesis
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3108
      unfolding compact_def
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3109
    proof safe
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3110
      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
  3111
      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
  3112
      then obtain K where
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3113
        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
  3114
        unfolding choice_iff by blast
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3115
      {
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3116
        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
  3117
        assume "e > 0"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3118
        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
  3119
          by simp_all
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3120
        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
  3121
        proof (rule ccontr)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3122
          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
  3123
            using `s \<noteq> {}`
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3124
            by auto
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3125
          moreover
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3126
          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
  3127
          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
  3128
          ultimately
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3129
          show False using f'
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3130
          proof (induct arbitrary: s f f' rule: finite_ne_induct)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3131
            case (singleton x)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3132
            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
  3133
            thus ?case using singleton by (auto simp: ball_def)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3134
          next
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3135
            case (insert x A)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3136
            show ?case
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3137
            proof cases
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3138
              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
  3139
              have "infinite ((f o f') -` \<Union>((\<lambda>x. ball x e) ` (insert x A)))"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3140
                using insert by (intro infinite_super[OF _ inf_ms]) auto
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3141
              also have "((f o f') -` \<Union>((\<lambda>x. ball x e) ` (insert x A))) =
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3142
                {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
  3143
              finally have "infinite \<dots>" .
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3144
              moreover assume "finite {m. (f o f') m \<in> ball x e}"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3145
              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
  3146
              hence "A \<noteq> {}" by auto then obtain k where "k \<in> A" by auto
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3147
              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
  3148
              have r_mono: "\<And>n m. n < m \<Longrightarrow> r n < r m"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3149
                using enumerate_mono[OF _ inf] by (simp add: r_def)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3150
              hence "subseq r" by (simp add: subseq_def)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3151
              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
  3152
                using enumerate_in_set[OF inf] by (simp add: r_def)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3153
              show False
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3154
              proof (rule insert)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3155
                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
  3156
                fix k s assume "k \<in> A" "subseq s"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3157
                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
  3158
                  by (subst (2) o_assoc[symmetric]) (intro insert(6) subseq_o, simp_all)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3159
              next
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3160
                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
  3161
              qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3162
            next
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3163
              assume inf: "infinite {m. (f o f') m \<in> ball x e}"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3164
              def r \<equiv> "enumerate {m. (f o f') m \<in> ball x e}"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3165
              have r_mono: "\<And>n m. n < m \<Longrightarrow> r n < r m"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3166
                using enumerate_mono[OF _ inf] by (simp add: r_def)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3167
              hence "subseq r" by (simp add: subseq_def)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3168
              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
  3169
              moreover
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3170
              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
  3171
                using enumerate_in_set[OF inf] by (simp add: r_def)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3172
              hence "(f o f') (r i) \<in> ball x e" by simp
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3173
              ultimately show False by simp
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3174
            qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3175
          qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3176
        qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3177
      }
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3178
      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
  3179
      let ?e = "\<lambda>n. 1 / real (Suc n)"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3180
      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
  3181
      interpret subseqs ?P using ex by unfold_locales force
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3182
      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
  3183
        by (simp add: complete_def)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3184
      have "\<exists>l\<in>s. (f o diagseq) ----> l"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3185
      proof (intro limI metric_CauchyI)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3186
        fix e::real assume "0 < e" hence "0 < e / 2" by auto
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3187
        from nat_approx_posE[OF this] guess n . note n = this
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3188
        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
  3189
        proof (rule exI[where x="Suc n"], safe)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3190
          fix m mm assume "Suc n \<le> m" "Suc n \<le> mm"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3191
          let ?e = "1 / real (Suc n)"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3192
          from reducer_reduces[of n] obtain k where
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3193
            "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
  3194
            unfolding seqseq_reducer by auto
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3195
          moreover
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3196
          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
  3197
          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
  3198
          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
  3199
          finally
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3200
          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
  3201
            by (intro add_strict_mono) auto
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3202
          hence "dist ((f \<circ> diagseq) m) k + dist ((f \<circ> diagseq) mm) k < e"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3203
            by (simp add: dist_commute)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3204
          moreover have "dist ((f \<circ> diagseq) m) ((f \<circ> diagseq) mm) \<le>
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3205
            dist ((f \<circ> diagseq) m) k + dist ((f \<circ> diagseq) mm) k"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3206
            by (rule dist_triangle2)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3207
          ultimately show "dist ((f \<circ> diagseq) m) ((f \<circ> diagseq) mm) < e"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3208
            by simp
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3209
        qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3210
      next
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3211
        fix n show "(f o diagseq) n \<in> s" using f by simp
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3212
      qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3213
      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
  3214
    qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3215
  qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3216
qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3217
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3218
lemma compact_eq_bounded_closed:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3219
  fixes s :: "'a::heine_borel set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3220
  shows "compact s \<longleftrightarrow> bounded s \<and> closed s"  (is "?lhs = ?rhs")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3221
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3222
  assume ?lhs thus ?rhs unfolding compact_eq_bolzano_weierstrass using bolzano_weierstrass_imp_bounded bolzano_weierstrass_imp_closed by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3223
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3224
  assume ?rhs thus ?lhs using bounded_closed_imp_compact by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3225
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3226
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3227
lemma compact_imp_bounded:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3228
  fixes s :: "'a::metric_space set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3229
  shows "compact s ==> bounded s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3230
proof -
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3231
  assume "compact s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3232
  hence "\<forall>f. (\<forall>t\<in>f. open t) \<and> s \<subseteq> \<Union>f \<longrightarrow> (\<exists>f'\<subseteq>f. finite f' \<and> s \<subseteq> \<Union>f')"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3233
    by (rule compact_imp_heine_borel)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3234
  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
  3235
    using heine_borel_imp_bolzano_weierstrass[of s] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3236
  thus "bounded s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3237
    by (rule bolzano_weierstrass_imp_bounded)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3238
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3239
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3240
lemma compact_imp_closed:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3241
  fixes s :: "'a::metric_space set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3242
  shows "compact s ==> closed s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3243
proof -
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3244
  assume "compact s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3245
  hence "\<forall>f. (\<forall>t\<in>f. open t) \<and> s \<subseteq> \<Union>f \<longrightarrow> (\<exists>f'\<subseteq>f. finite f' \<and> s \<subseteq> \<Union>f')"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3246
    by (rule compact_imp_heine_borel)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3247
  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
  3248
    using heine_borel_imp_bolzano_weierstrass[of s] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3249
  thus "closed s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3250
    by (rule bolzano_weierstrass_imp_closed)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3251
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3252
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3253
text{* In particular, some common special cases. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3254
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3255
lemma compact_empty[simp]:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3256
 "compact {}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3257
  unfolding compact_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3258
  by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3259
44075
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3260
lemma compact_union [intro]:
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3261
  assumes "compact s" and "compact t"
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3262
  shows "compact (s \<union> t)"
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3263
proof (rule compactI)
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3264
  fix f :: "nat \<Rightarrow> 'a"
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3265
  assume "\<forall>n. f n \<in> s \<union> t"
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3266
  hence "infinite {n. f n \<in> s \<union> t}" by simp
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3267
  hence "infinite {n. f n \<in> s} \<or> infinite {n. f n \<in> t}" by simp
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3268
  thus "\<exists>l\<in>s \<union> t. \<exists>r. subseq r \<and> ((f \<circ> r) ---> l) sequentially"
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3269
  proof
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3270
    assume "infinite {n. f n \<in> s}"
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3271
    from infinite_enumerate [OF this]
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3272
    obtain q where "subseq q" "\<forall>n. (f \<circ> q) n \<in> s" by auto
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3273
    obtain r l where "l \<in> s" "subseq r" "((f \<circ> q \<circ> r) ---> l) sequentially"
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3274
      using `compact s` `\<forall>n. (f \<circ> q) n \<in> s` by (rule compactE)
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3275
    hence "l \<in> s \<union> t" "subseq (q \<circ> r)" "((f \<circ> (q \<circ> r)) ---> l) sequentially"
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3276
      using `subseq q` by (simp_all add: subseq_o o_assoc)
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3277
    thus ?thesis by auto
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3278
  next
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3279
    assume "infinite {n. f n \<in> t}"
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3280
    from infinite_enumerate [OF this]
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3281
    obtain q where "subseq q" "\<forall>n. (f \<circ> q) n \<in> t" by auto
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3282
    obtain r l where "l \<in> t" "subseq r" "((f \<circ> q \<circ> r) ---> l) sequentially"
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3283
      using `compact t` `\<forall>n. (f \<circ> q) n \<in> t` by (rule compactE)
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3284
    hence "l \<in> s \<union> t" "subseq (q \<circ> r)" "((f \<circ> (q \<circ> r)) ---> l) sequentially"
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3285
      using `subseq q` by (simp_all add: subseq_o o_assoc)
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3286
    thus ?thesis by auto
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3287
  qed
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3288
qed
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3289
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3290
lemma compact_Union [intro]: "finite S \<Longrightarrow> (\<And>T. T \<in> S \<Longrightarrow> compact T) \<Longrightarrow> compact (\<Union>S)"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3291
  by (induct set: finite) auto
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3292
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3293
lemma compact_UN [intro]:
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3294
  "finite A \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> compact (B x)) \<Longrightarrow> compact (\<Union>x\<in>A. B x)"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3295
  unfolding SUP_def by (rule compact_Union) auto
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3296
44075
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3297
lemma compact_inter_closed [intro]:
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3298
  assumes "compact s" and "closed t"
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3299
  shows "compact (s \<inter> t)"
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3300
proof (rule compactI)
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3301
  fix f :: "nat \<Rightarrow> 'a"
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3302
  assume "\<forall>n. f n \<in> s \<inter> t"
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3303
  hence "\<forall>n. f n \<in> s" and "\<forall>n. f n \<in> t" by simp_all
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3304
  obtain l r where "l \<in> s" "subseq r" "((f \<circ> r) ---> l) sequentially"
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3305
    using `compact s` `\<forall>n. f n \<in> s` by (rule compactE)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3306
  moreover
44075
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3307
  from `closed t` `\<forall>n. f n \<in> t` `((f \<circ> r) ---> l) sequentially` have "l \<in> t"
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3308
    unfolding closed_sequential_limits o_def by fast
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3309
  ultimately show "\<exists>l\<in>s \<inter> t. \<exists>r. subseq r \<and> ((f \<circ> r) ---> l) sequentially"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3310
    by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3311
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3312
44075
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3313
lemma closed_inter_compact [intro]:
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3314
  assumes "closed s" and "compact t"
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3315
  shows "compact (s \<inter> t)"
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3316
  using compact_inter_closed [of t s] assms
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3317
  by (simp add: Int_commute)
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3318
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3319
lemma compact_inter [intro]:
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3320
  assumes "compact s" and "compact t"
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3321
  shows "compact (s \<inter> t)"
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3322
  using assms by (intro compact_inter_closed compact_imp_closed)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3323
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3324
lemma compact_sing [simp]: "compact {a}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3325
  unfolding compact_def o_def subseq_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3326
  by (auto simp add: tendsto_const)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3327
44075
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3328
lemma compact_insert [simp]:
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3329
  assumes "compact s" shows "compact (insert x s)"
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3330
proof -
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3331
  have "compact ({x} \<union> s)"
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3332
    using compact_sing assms by (rule compact_union)
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3333
  thus ?thesis by simp
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3334
qed
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3335
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3336
lemma finite_imp_compact:
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3337
  shows "finite s \<Longrightarrow> compact s"
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3338
  by (induct set: finite) simp_all
5952bd355779 generalize more lemmas about compactness
huffman
parents: 44074
diff changeset
  3339
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3340
lemma compact_cball[simp]:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3341
  fixes x :: "'a::heine_borel"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3342
  shows "compact(cball x e)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3343
  using compact_eq_bounded_closed bounded_cball closed_cball
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3344
  by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3345
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3346
lemma compact_frontier_bounded[intro]:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3347
  fixes s :: "'a::heine_borel set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3348
  shows "bounded s ==> compact(frontier s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3349
  unfolding frontier_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3350
  using compact_eq_bounded_closed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3351
  by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3352
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3353
lemma compact_frontier[intro]:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3354
  fixes s :: "'a::heine_borel set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3355
  shows "compact s ==> compact (frontier s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3356
  using compact_eq_bounded_closed compact_frontier_bounded
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3357
  by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3358
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3359
lemma frontier_subset_compact:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3360
  fixes s :: "'a::heine_borel set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3361
  shows "compact s ==> frontier s \<subseteq> s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3362
  using frontier_subset_closed compact_eq_bounded_closed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3363
  by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3364
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3365
lemma open_delete:
36668
941ba2da372e simplify definition of t1_space; generalize lemma closed_sing and related lemmas
huffman
parents: 36667
diff changeset
  3366
  fixes s :: "'a::t1_space set"
941ba2da372e simplify definition of t1_space; generalize lemma closed_sing and related lemmas
huffman
parents: 36667
diff changeset
  3367
  shows "open s \<Longrightarrow> open (s - {x})"
941ba2da372e simplify definition of t1_space; generalize lemma closed_sing and related lemmas
huffman
parents: 36667
diff changeset
  3368
  by (simp add: open_Diff)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3369
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3370
text{* Finite intersection property. I could make it an equivalence in fact. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3371
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3372
lemma compact_imp_fip:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3373
  assumes "compact s"  "\<forall>t \<in> f. closed t"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3374
        "\<forall>f'. finite f' \<and> f' \<subseteq> f --> (s \<inter> (\<Inter> f') \<noteq> {})"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3375
  shows "s \<inter> (\<Inter> f) \<noteq> {}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3376
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3377
  assume as:"s \<inter> (\<Inter> f) = {}"
34105
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
  3378
  hence "s \<subseteq> \<Union> uminus ` f" by auto
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
  3379
  moreover have "Ball (uminus ` f) open" using open_Diff closed_Diff using assms(2) by auto
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
  3380
  ultimately obtain f' where f':"f' \<subseteq> uminus ` f"  "finite f'"  "s \<subseteq> \<Union>f'" using assms(1)[unfolded compact_eq_heine_borel, THEN spec[where x="(\<lambda>t. - t) ` f"]] by auto
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
  3381
  hence "finite (uminus ` f') \<and> uminus ` f' \<subseteq> f" by(auto simp add: Diff_Diff_Int)
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
  3382
  hence "s \<inter> \<Inter>uminus ` f' \<noteq> {}" using assms(3)[THEN spec[where x="uminus ` f'"]] by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3383
  thus False using f'(3) unfolding subset_eq and Union_iff by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3384
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3385
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  3386
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  3387
subsection {* Bounded closed nest property (proof does not use Heine-Borel) *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3388
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3389
lemma bounded_closed_nest:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3390
  assumes "\<forall>n. closed(s n)" "\<forall>n. (s n \<noteq> {})"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3391
  "(\<forall>m n. m \<le> n --> s n \<subseteq> s m)"  "bounded(s 0)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3392
  shows "\<exists>a::'a::heine_borel. \<forall>n::nat. a \<in> s(n)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3393
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3394
  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
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3395
  from assms(4,1) have *:"compact (s 0)" using bounded_closed_imp_compact[of "s 0"] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3396
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3397
  then obtain l r where lr:"l\<in>s 0" "subseq r" "((x \<circ> r) ---> l) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3398
    unfolding compact_def apply(erule_tac x=x in allE)  using x using assms(3) by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3399
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3400
  { fix n::nat
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3401
    { fix e::real assume "e>0"
44907
93943da0a010 remove redundant lemma Lim_sequentially in favor of lemma LIMSEQ_def
huffman
parents: 44905
diff changeset
  3402
      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
  3403
      hence "dist ((x \<circ> r) (max N n)) l < e" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3404
      moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3405
      have "r (max N n) \<ge> n" using lr(2) using subseq_bigger[of r "max N n"] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3406
      hence "(x \<circ> r) (max N n) \<in> s n"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3407
        using x apply(erule_tac x=n in allE)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3408
        using x apply(erule_tac x="r (max N n)" in allE)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3409
        using assms(3) apply(erule_tac x=n in allE)apply( erule_tac x="r (max N n)" in allE) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3410
      ultimately have "\<exists>y\<in>s n. dist y l < e" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3411
    }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3412
    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
  3413
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3414
  thus ?thesis by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3415
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3416
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  3417
text {* Decreasing case does not even need compactness, just completeness. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3418
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3419
lemma decreasing_closed_nest:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3420
  assumes "\<forall>n. closed(s n)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3421
          "\<forall>n. (s n \<noteq> {})"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3422
          "\<forall>m n. m \<le> n --> s n \<subseteq> s m"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3423
          "\<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
  3424
  shows "\<exists>a::'a::complete_space. \<forall>n::nat. a \<in> s n"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3425
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3426
  have "\<forall>n. \<exists> x. x\<in>s n" using assms(2) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3427
  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
  3428
  then obtain t where t: "\<forall>n. t n \<in> s n" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3429
  { fix e::real assume "e>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3430
    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
  3431
    { fix m n ::nat assume "N \<le> m \<and> N \<le> n"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3432
      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
  3433
      hence "dist (t m) (t n) < e" using N by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3434
    }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3435
    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
  3436
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3437
  hence  "Cauchy t" unfolding cauchy_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3438
  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
  3439
  { fix n::nat
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3440
    { fix e::real assume "e>0"
44907
93943da0a010 remove redundant lemma Lim_sequentially in favor of lemma LIMSEQ_def
huffman
parents: 44905
diff changeset
  3441
      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
  3442
      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
  3443
      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
  3444
    }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3445
    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
  3446
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3447
  then show ?thesis by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3448
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3449
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  3450
text {* Strengthen it to the intersection actually being a singleton. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3451
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3452
lemma decreasing_closed_nest_sing:
44632
076a45f65e12 simplify/generalize some proofs
huffman
parents: 44628
diff changeset
  3453
  fixes s :: "nat \<Rightarrow> 'a::complete_space set"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3454
  assumes "\<forall>n. closed(s n)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3455
          "\<forall>n. s n \<noteq> {}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3456
          "\<forall>m n. m \<le> n --> s n \<subseteq> s m"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3457
          "\<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
  3458
  shows "\<exists>a. \<Inter>(range s) = {a}"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3459
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3460
  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
  3461
  { fix b assume b:"b \<in> \<Inter>(range s)"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3462
    { fix e::real assume "e>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3463
      hence "dist a b < e" using assms(4 )using b using a by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3464
    }
36778
739a9379e29b avoid using real-specific versions of generic lemmas
huffman
parents: 36669
diff changeset
  3465
    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
  3466
  }
34104
22758f95e624 re-state lemmas using 'range'
huffman
parents: 33758
diff changeset
  3467
  with a have "\<Inter>(range s) = {a}" unfolding image_def by auto
22758f95e624 re-state lemmas using 'range'
huffman
parents: 33758
diff changeset
  3468
  thus ?thesis ..
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3469
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3470
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3471
text{* Cauchy-type criteria for uniform convergence. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3472
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3473
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
  3474
 "(\<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
  3475
  (\<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
  3476
proof(rule)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3477
  assume ?lhs
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3478
  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
  3479
  { fix e::real assume "e>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3480
    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
  3481
    { 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
  3482
      hence "dist (s m x) (s n x) < e"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3483
        using N[THEN spec[where x=m], THEN spec[where x=x]]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3484
        using N[THEN spec[where x=n], THEN spec[where x=x]]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3485
        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
  3486
    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
  3487
  thus ?rhs 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
  assume ?rhs
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3490
  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
  3491
  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
  3492
    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
  3493
  { fix e::real assume "e>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3494
    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
  3495
      using `?rhs`[THEN spec[where x="e/2"]] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3496
    { fix x assume "P x"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3497
      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
  3498
        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
  3499
      fix n::nat assume "n\<ge>N"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3500
      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
  3501
        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
  3502
    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
  3503
  thus ?lhs by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3504
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3505
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3506
lemma uniformly_cauchy_imp_uniformly_convergent:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3507
  fixes s :: "nat \<Rightarrow> 'a \<Rightarrow> 'b::heine_borel"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3508
  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
  3509
          "\<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
  3510
  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
  3511
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3512
  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
  3513
    using assms(1) unfolding uniformly_convergent_eq_cauchy[THEN sym] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3514
  moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3515
  { fix x assume "P x"
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41969
diff changeset
  3516
    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
  3517
      using l and assms(2) unfolding LIMSEQ_def by blast  }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3518
  ultimately show ?thesis by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3519
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3520
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  3521
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  3522
subsection {* Continuity *}
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  3523
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  3524
text {* Define continuity over a net to take in restrictions of the set. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3525
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3526
definition
44081
730f7cced3a6 rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents: 44076
diff changeset
  3527
  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
  3528
  where "continuous net f \<longleftrightarrow> (f ---> f(netlimit net)) net"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3529
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3530
lemma continuous_trivial_limit:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3531
 "trivial_limit net ==> continuous net f"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3532
  unfolding continuous_def tendsto_def trivial_limit_eq by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3533
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3534
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
  3535
  unfolding continuous_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3536
  unfolding tendsto_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3537
  using netlimit_within[of x s]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3538
  by (cases "trivial_limit (at x within s)") (auto simp add: trivial_limit_eventually)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3539
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3540
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
  3541
  using continuous_within [of x UNIV f] by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3542
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3543
lemma continuous_at_within:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3544
  assumes "continuous (at x) f"  shows "continuous (at x within s) f"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3545
  using assms unfolding continuous_at continuous_within
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3546
  by (rule Lim_at_within)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3547
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3548
text{* Derive the epsilon-delta forms, which we often use as "definitions" *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3549
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3550
lemma continuous_within_eps_delta:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3551
  "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
  3552
  unfolding continuous_within and Lim_within
44584
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  3553
  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
  3554
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3555
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
  3556
                           \<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
  3557
  using continuous_within_eps_delta [of x UNIV f] by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3558
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3559
text{* Versions in terms of open balls. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3560
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3561
lemma continuous_within_ball:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3562
 "continuous (at x within s) f \<longleftrightarrow> (\<forall>e>0. \<exists>d>0.
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3563
                            f ` (ball x d \<inter> s) \<subseteq> ball (f x) e)" (is "?lhs = ?rhs")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3564
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3565
  assume ?lhs
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3566
  { fix e::real assume "e>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3567
    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
  3568
      using `?lhs`[unfolded continuous_within Lim_within] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3569
    { fix y assume "y\<in>f ` (ball x d \<inter> s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3570
      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
  3571
        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
  3572
    }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3573
    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
  3574
  thus ?rhs by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3575
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3576
  assume ?rhs thus ?lhs unfolding continuous_within Lim_within ball_def subset_eq
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3577
    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
  3578
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3579
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3580
lemma continuous_at_ball:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3581
  "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
  3582
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3583
  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
  3584
    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
  3585
    unfolding dist_nz[THEN sym] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3586
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3587
  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
  3588
    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
  3589
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3590
36440
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3591
text{* Define setwise continuity in terms of limits within the set. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3592
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3593
definition
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  3594
  continuous_on ::
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  3595
    "'a set \<Rightarrow> ('a::topological_space \<Rightarrow> 'b::topological_space) \<Rightarrow> bool"
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  3596
where
36440
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3597
  "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
  3598
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3599
lemma continuous_on_topological:
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  3600
  "continuous_on s f \<longleftrightarrow>
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  3601
    (\<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
  3602
      (\<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
  3603
unfolding continuous_on_def tendsto_def
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3604
unfolding Limits.eventually_within eventually_at_topological
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3605
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
  3606
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  3607
lemma continuous_on_iff:
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  3608
  "continuous_on s f \<longleftrightarrow>
36440
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3609
    (\<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
  3610
unfolding continuous_on_def Lim_within
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3611
apply (intro ball_cong [OF refl] all_cong ex_cong)
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3612
apply (rename_tac y, case_tac "y = x", simp)
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3613
apply (simp add: dist_nz)
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  3614
done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3615
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3616
definition
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3617
  uniformly_continuous_on ::
36440
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3618
    "'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
  3619
where
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3620
  "uniformly_continuous_on s f \<longleftrightarrow>
36440
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3621
    (\<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
  3622
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3623
text{* Some simple consequential lemmas. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3624
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3625
lemma uniformly_continuous_imp_continuous:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3626
 " uniformly_continuous_on s f ==> continuous_on s f"
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  3627
  unfolding uniformly_continuous_on_def continuous_on_iff by blast
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3628
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3629
lemma continuous_at_imp_continuous_within:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3630
 "continuous (at x) f ==> continuous (at x within s) f"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3631
  unfolding continuous_within continuous_at using Lim_at_within by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3632
36440
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3633
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
  3634
unfolding tendsto_def by (simp add: trivial_limit_eq)
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3635
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  3636
lemma continuous_at_imp_continuous_on:
36440
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3637
  assumes "\<forall>x\<in>s. continuous (at x) f"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3638
  shows "continuous_on s f"
36440
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3639
unfolding continuous_on_def
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3640
proof
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3641
  fix x assume "x \<in> s"
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3642
  with assms have *: "(f ---> f (netlimit (at x))) (at x)"
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3643
    unfolding continuous_def by simp
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3644
  have "(f ---> f x) (at x)"
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3645
  proof (cases "trivial_limit (at x)")
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3646
    case True thus ?thesis
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3647
      by (rule Lim_trivial_limit)
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3648
  next
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3649
    case False
36667
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  3650
    hence 1: "netlimit (at x) = x"
45031
9583f2b56f85 add lemmas within_empty and tendsto_bot;
huffman
parents: 44909
diff changeset
  3651
      using netlimit_within [of x UNIV] by simp
36440
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3652
    with * show ?thesis by simp
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3653
  qed
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3654
  thus "(f ---> f x) (at x within s)"
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3655
    by (rule Lim_at_within)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3656
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3657
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3658
lemma continuous_on_eq_continuous_within:
36440
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3659
  "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
  3660
unfolding continuous_on_def continuous_def
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3661
apply (rule ball_cong [OF refl])
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3662
apply (case_tac "trivial_limit (at x within s)")
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3663
apply (simp add: Lim_trivial_limit)
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3664
apply (simp add: netlimit_within)
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3665
done
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3666
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3667
lemmas continuous_on = continuous_on_def -- "legacy theorem name"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3668
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3669
lemma continuous_on_eq_continuous_at:
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  3670
  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
  3671
  by (auto simp add: continuous_on continuous_at Lim_within_open)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3672
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3673
lemma continuous_within_subset:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3674
 "continuous (at x within s) f \<Longrightarrow> t \<subseteq> s
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3675
             ==> continuous (at x within t) f"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3676
  unfolding continuous_within by(metis Lim_within_subset)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3677
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3678
lemma continuous_on_subset:
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  3679
  shows "continuous_on s f \<Longrightarrow> t \<subseteq> s ==> continuous_on t f"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3680
  unfolding continuous_on by (metis subset_eq Lim_within_subset)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3681
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3682
lemma continuous_on_interior:
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  3683
  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
  3684
  by (erule interiorE, drule (1) continuous_on_subset,
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  3685
    simp add: continuous_on_eq_continuous_at)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3686
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3687
lemma continuous_on_eq:
36440
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3688
  "(\<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
  3689
  unfolding continuous_on_def tendsto_def Limits.eventually_within
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3690
  by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3691
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  3692
text {* Characterization of various kinds of continuity in terms of sequences. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3693
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3694
lemma continuous_within_sequentially:
44533
7abe4a59f75d generalize and simplify proof of continuous_within_sequentially
huffman
parents: 44531
diff changeset
  3695
  fixes f :: "'a::metric_space \<Rightarrow> 'b::topological_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3696
  shows "continuous (at a within s) f \<longleftrightarrow>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3697
                (\<forall>x. (\<forall>n::nat. x n \<in> s) \<and> (x ---> a) sequentially
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3698
                     --> ((f o x) ---> f a) sequentially)" (is "?lhs = ?rhs")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3699
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3700
  assume ?lhs
44533
7abe4a59f75d generalize and simplify proof of continuous_within_sequentially
huffman
parents: 44531
diff changeset
  3701
  { 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
  3702
    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
  3703
    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
  3704
      unfolding continuous_within tendsto_def eventually_within by auto
7abe4a59f75d generalize and simplify proof of continuous_within_sequentially
huffman
parents: 44531
diff changeset
  3705
    have "eventually (\<lambda>n. dist (x n) a < d) sequentially"
7abe4a59f75d generalize and simplify proof of continuous_within_sequentially
huffman
parents: 44531
diff changeset
  3706
      using x(2) `d>0` by simp
7abe4a59f75d generalize and simplify proof of continuous_within_sequentially
huffman
parents: 44531
diff changeset
  3707
    hence "eventually (\<lambda>n. (f \<circ> x) n \<in> T) sequentially"
46887
cb891d9a23c1 use eventually_elim method
noschinl
parents: 45776
diff changeset
  3708
    proof eventually_elim
cb891d9a23c1 use eventually_elim method
noschinl
parents: 45776
diff changeset
  3709
      case (elim n) thus ?case
44533
7abe4a59f75d generalize and simplify proof of continuous_within_sequentially
huffman
parents: 44531
diff changeset
  3710
        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
  3711
    qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3712
  }
44533
7abe4a59f75d generalize and simplify proof of continuous_within_sequentially
huffman
parents: 44531
diff changeset
  3713
  thus ?rhs unfolding tendsto_iff unfolding tendsto_def by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3714
next
44533
7abe4a59f75d generalize and simplify proof of continuous_within_sequentially
huffman
parents: 44531
diff changeset
  3715
  assume ?rhs thus ?lhs
7abe4a59f75d generalize and simplify proof of continuous_within_sequentially
huffman
parents: 44531
diff changeset
  3716
    unfolding continuous_within tendsto_def [where l="f a"]
7abe4a59f75d generalize and simplify proof of continuous_within_sequentially
huffman
parents: 44531
diff changeset
  3717
    by (simp add: sequentially_imp_eventually_within)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3718
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3719
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3720
lemma continuous_at_sequentially:
44533
7abe4a59f75d generalize and simplify proof of continuous_within_sequentially
huffman
parents: 44531
diff changeset
  3721
  fixes f :: "'a::metric_space \<Rightarrow> 'b::topological_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3722
  shows "continuous (at a) f \<longleftrightarrow> (\<forall>x. (x ---> a) sequentially
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3723
                  --> ((f o x) ---> f a) sequentially)"
45031
9583f2b56f85 add lemmas within_empty and tendsto_bot;
huffman
parents: 44909
diff changeset
  3724
  using continuous_within_sequentially[of a UNIV f] by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3725
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3726
lemma continuous_on_sequentially:
44533
7abe4a59f75d generalize and simplify proof of continuous_within_sequentially
huffman
parents: 44531
diff changeset
  3727
  fixes f :: "'a::metric_space \<Rightarrow> 'b::topological_space"
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  3728
  shows "continuous_on s f \<longleftrightarrow>
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  3729
    (\<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
  3730
                    --> ((f o x) ---> f(a)) sequentially)" (is "?lhs = ?rhs")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3731
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3732
  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
  3733
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3734
  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
  3735
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3736
44648
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  3737
lemma uniformly_continuous_on_sequentially:
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  3738
  "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
  3739
                    ((\<lambda>n. dist (x n) (y n)) ---> 0) sequentially
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  3740
                    \<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
  3741
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3742
  assume ?lhs
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  3743
  { 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
  3744
    { fix e::real assume "e>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3745
      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
  3746
        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
  3747
      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
  3748
      { fix n assume "n\<ge>N"
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  3749
        hence "dist (f (x n)) (f (y n)) < e"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3750
          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
  3751
          unfolding dist_commute by simp  }
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  3752
      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
  3753
    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
  3754
  thus ?rhs by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3755
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3756
  assume ?rhs
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3757
  { assume "\<not> ?lhs"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3758
    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
  3759
    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
  3760
      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
  3761
      by (auto simp add: dist_commute)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3762
    def x \<equiv> "\<lambda>n::nat. fst (fa (inverse (real n + 1)))"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3763
    def y \<equiv> "\<lambda>n::nat. snd (fa (inverse (real n + 1)))"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3764
    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
  3765
      unfolding x_def and y_def using fa by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3766
    { fix e::real assume "e>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3767
      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
  3768
      { fix n::nat assume "n\<ge>N"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3769
        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
  3770
        also have "\<dots> < e" using N by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3771
        finally have "inverse (real n + 1) < e" by auto
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  3772
        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
  3773
      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
  3774
    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
  3775
    hence False using fxy and `e>0` by auto  }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3776
  thus ?lhs unfolding uniformly_continuous_on_def by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3777
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3778
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3779
text{* The usual transformation theorems. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3780
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3781
lemma continuous_transform_within:
36667
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  3782
  fixes f g :: "'a::metric_space \<Rightarrow> 'b::topological_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3783
  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
  3784
          "continuous (at x within s) f"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3785
  shows "continuous (at x within s) g"
36667
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  3786
unfolding continuous_within
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  3787
proof (rule Lim_transform_within)
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  3788
  show "0 < d" by fact
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  3789
  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
  3790
    using assms(3) by auto
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  3791
  have "f x = g x"
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  3792
    using assms(1,2,3) by auto
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  3793
  thus "(f ---> g x) (at x within s)"
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  3794
    using assms(4) unfolding continuous_within by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3795
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3796
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3797
lemma continuous_transform_at:
36667
21404f7dec59 generalize some lemmas
huffman
parents: 36660
diff changeset
  3798
  fixes f g :: "'a::metric_space \<Rightarrow> 'b::topological_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3799
  assumes "0 < d" "\<forall>x'. dist x' x < d --> f x' = g x'"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3800
          "continuous (at x) f"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3801
  shows "continuous (at x) g"
45031
9583f2b56f85 add lemmas within_empty and tendsto_bot;
huffman
parents: 44909
diff changeset
  3802
  using continuous_transform_within [of d x UNIV f g] assms by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3803
44648
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  3804
subsubsection {* Structural rules for pointwise continuity *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3805
44647
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3806
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
  3807
  unfolding continuous_within by (rule tendsto_ident_at_within)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3808
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3809
lemma continuous_at_id: "continuous (at a) (\<lambda>x. x)"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3810
  unfolding continuous_at by (rule tendsto_ident_at)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3811
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3812
lemma continuous_const: "continuous F (\<lambda>x. c)"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3813
  unfolding continuous_def by (rule tendsto_const)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3814
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3815
lemma continuous_dist:
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3816
  assumes "continuous F f" and "continuous F g"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3817
  shows "continuous F (\<lambda>x. dist (f x) (g x))"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3818
  using assms unfolding continuous_def by (rule tendsto_dist)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3819
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3820
lemma continuous_infdist:
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3821
  assumes "continuous F f"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3822
  shows "continuous F (\<lambda>x. infdist (f x) A)"
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3823
  using assms unfolding continuous_def by (rule tendsto_infdist)
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  3824
44647
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3825
lemma continuous_norm:
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3826
  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
  3827
  unfolding continuous_def by (rule tendsto_norm)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3828
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3829
lemma continuous_infnorm:
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3830
  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
  3831
  unfolding continuous_def by (rule tendsto_infnorm)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3832
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3833
lemma continuous_add:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3834
  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
  3835
  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
  3836
  unfolding continuous_def by (rule tendsto_add)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3837
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3838
lemma continuous_minus:
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3839
  fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3840
  shows "continuous F f \<Longrightarrow> continuous F (\<lambda>x. - f x)"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3841
  unfolding continuous_def by (rule tendsto_minus)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3842
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3843
lemma continuous_diff:
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3844
  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
  3845
  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
  3846
  unfolding continuous_def by (rule tendsto_diff)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3847
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3848
lemma continuous_scaleR:
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3849
  fixes g :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3850
  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
  3851
  unfolding continuous_def by (rule tendsto_scaleR)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3852
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3853
lemma continuous_mult:
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3854
  fixes f g :: "'a::t2_space \<Rightarrow> 'b::real_normed_algebra"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3855
  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
  3856
  unfolding continuous_def by (rule tendsto_mult)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3857
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3858
lemma continuous_inner:
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3859
  assumes "continuous F f" and "continuous F g"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3860
  shows "continuous F (\<lambda>x. inner (f x) (g x))"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3861
  using assms unfolding continuous_def by (rule tendsto_inner)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3862
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3863
lemma continuous_inverse:
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3864
  fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_div_algebra"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3865
  assumes "continuous F f" and "f (netlimit F) \<noteq> 0"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3866
  shows "continuous F (\<lambda>x. inverse (f x))"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3867
  using assms unfolding continuous_def by (rule tendsto_inverse)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3868
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3869
lemma continuous_at_within_inverse:
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3870
  fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_div_algebra"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3871
  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
  3872
  shows "continuous (at a within s) (\<lambda>x. inverse (f x))"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3873
  using assms unfolding continuous_within by (rule tendsto_inverse)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3874
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3875
lemma continuous_at_inverse:
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3876
  fixes f :: "'a::t2_space \<Rightarrow> 'b::real_normed_div_algebra"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3877
  assumes "continuous (at a) f" and "f a \<noteq> 0"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3878
  shows "continuous (at a) (\<lambda>x. inverse (f x))"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3879
  using assms unfolding continuous_at by (rule tendsto_inverse)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3880
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3881
lemmas continuous_intros = continuous_at_id continuous_within_id
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3882
  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
  3883
  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
  3884
  continuous_inner continuous_at_inverse continuous_at_within_inverse
34964
4e8be3c04d37 Replaced vec1 and dest_vec1 by abbreviation.
hoelzl
parents: 34291
diff changeset
  3885
44648
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  3886
subsubsection {* Structural rules for setwise continuity *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3887
44647
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3888
lemma continuous_on_id: "continuous_on s (\<lambda>x. x)"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3889
  unfolding continuous_on_def by (fast intro: tendsto_ident_at_within)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3890
44531
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  3891
lemma continuous_on_const: "continuous_on s (\<lambda>x. c)"
44125
230a8665c919 mark some redundant theorems as legacy
huffman
parents: 44122
diff changeset
  3892
  unfolding continuous_on_def by (auto intro: tendsto_intros)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3893
44647
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3894
lemma continuous_on_norm:
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3895
  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
  3896
  unfolding continuous_on_def by (fast intro: tendsto_norm)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3897
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3898
lemma continuous_on_infnorm:
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3899
  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
  3900
  unfolding continuous_on by (fast intro: tendsto_infnorm)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3901
44531
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  3902
lemma continuous_on_minus:
36440
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3903
  fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3904
  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
  3905
  unfolding continuous_on_def by (auto intro: tendsto_intros)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3906
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3907
lemma continuous_on_add:
36440
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3908
  fixes f g :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3909
  shows "continuous_on s f \<Longrightarrow> continuous_on s g
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3910
           \<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
  3911
  unfolding continuous_on_def by (auto intro: tendsto_intros)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3912
44531
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  3913
lemma continuous_on_diff:
36440
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  3914
  fixes f g :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3915
  shows "continuous_on s f \<Longrightarrow> continuous_on s g
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3916
           \<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
  3917
  unfolding continuous_on_def by (auto intro: tendsto_intros)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3918
44531
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  3919
lemma (in bounded_linear) continuous_on:
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  3920
  "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
  3921
  unfolding continuous_on_def by (fast intro: tendsto)
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  3922
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  3923
lemma (in bounded_bilinear) continuous_on:
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  3924
  "\<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
  3925
  unfolding continuous_on_def by (fast intro: tendsto)
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  3926
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  3927
lemma continuous_on_scaleR:
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  3928
  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
  3929
  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
  3930
  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
  3931
  using bounded_bilinear_scaleR assms
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  3932
  by (rule bounded_bilinear.continuous_on)
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  3933
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  3934
lemma continuous_on_mult:
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  3935
  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
  3936
  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
  3937
  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
  3938
  using bounded_bilinear_mult assms
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  3939
  by (rule bounded_bilinear.continuous_on)
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  3940
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  3941
lemma continuous_on_inner:
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  3942
  fixes g :: "'a::topological_space \<Rightarrow> 'b::real_inner"
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  3943
  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
  3944
  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
  3945
  using bounded_bilinear_inner assms
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  3946
  by (rule bounded_bilinear.continuous_on)
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44530
diff changeset
  3947
44647
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3948
lemma continuous_on_inverse:
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3949
  fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_div_algebra"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3950
  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
  3951
  shows "continuous_on s (\<lambda>x. inverse (f x))"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3952
  using assms unfolding continuous_on by (fast intro: tendsto_inverse)
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3953
44648
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  3954
subsubsection {* Structural rules for uniform continuity *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3955
44647
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3956
lemma uniformly_continuous_on_id:
44648
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  3957
  shows "uniformly_continuous_on s (\<lambda>x. x)"
44647
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3958
  unfolding uniformly_continuous_on_def by auto
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  3959
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3960
lemma uniformly_continuous_on_const:
44648
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  3961
  shows "uniformly_continuous_on s (\<lambda>x. c)"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3962
  unfolding uniformly_continuous_on_def by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3963
44648
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  3964
lemma uniformly_continuous_on_dist:
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  3965
  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
  3966
  assumes "uniformly_continuous_on s f"
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  3967
  assumes "uniformly_continuous_on s g"
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  3968
  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
  3969
proof -
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  3970
  { 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
  3971
      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
  3972
      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
  3973
      by arith
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  3974
  } note le = this
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  3975
  { fix x y
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  3976
    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
  3977
    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
  3978
    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
  3979
      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
  3980
        simp add: le)
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  3981
  }
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  3982
  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
  3983
    unfolding dist_real_def by simp
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  3984
qed
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  3985
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  3986
lemma uniformly_continuous_on_norm:
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  3987
  assumes "uniformly_continuous_on s f"
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  3988
  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
  3989
  unfolding norm_conv_dist using assms
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  3990
  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
  3991
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  3992
lemma (in bounded_linear) uniformly_continuous_on:
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  3993
  assumes "uniformly_continuous_on s g"
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  3994
  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
  3995
  using assms unfolding uniformly_continuous_on_sequentially
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  3996
  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
  3997
  by (auto intro: tendsto_zero)
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  3998
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3999
lemma uniformly_continuous_on_cmul:
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4000
  fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4001
  assumes "uniformly_continuous_on s f"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4002
  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
  4003
  using bounded_linear_scaleR_right assms
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4004
  by (rule bounded_linear.uniformly_continuous_on)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4005
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4006
lemma dist_minus:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4007
  fixes x y :: "'a::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4008
  shows "dist (- x) (- y) = dist x y"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4009
  unfolding dist_norm minus_diff_minus norm_minus_cancel ..
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4010
44648
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4011
lemma uniformly_continuous_on_minus:
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4012
  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
  4013
  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
  4014
  unfolding uniformly_continuous_on_def dist_minus .
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4015
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4016
lemma uniformly_continuous_on_add:
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4017
  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
  4018
  assumes "uniformly_continuous_on s f"
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4019
  assumes "uniformly_continuous_on s g"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4020
  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
  4021
  using assms unfolding uniformly_continuous_on_sequentially
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4022
  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
  4023
  by (auto intro: tendsto_add_zero)
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4024
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4025
lemma uniformly_continuous_on_diff:
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4026
  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
  4027
  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
  4028
  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
  4029
  unfolding ab_diff_minus using assms
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4030
  by (intro uniformly_continuous_on_add uniformly_continuous_on_minus)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4031
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4032
text{* Continuity of all kinds is preserved under composition. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4033
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4034
lemma continuous_within_topological:
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4035
  "continuous (at x within s) f \<longleftrightarrow>
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4036
    (\<forall>B. open B \<longrightarrow> f x \<in> B \<longrightarrow>
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4037
      (\<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
  4038
unfolding continuous_within
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4039
unfolding tendsto_def Limits.eventually_within eventually_at_topological
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4040
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
  4041
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4042
lemma continuous_within_compose:
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4043
  assumes "continuous (at x within s) f"
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4044
  assumes "continuous (at (f x) within f ` s) g"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4045
  shows "continuous (at x within s) (g o f)"
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4046
using assms unfolding continuous_within_topological by simp metis
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4047
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4048
lemma continuous_at_compose:
45031
9583f2b56f85 add lemmas within_empty and tendsto_bot;
huffman
parents: 44909
diff changeset
  4049
  assumes "continuous (at x) f" and "continuous (at (f x)) g"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4050
  shows "continuous (at x) (g o f)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4051
proof-
45031
9583f2b56f85 add lemmas within_empty and tendsto_bot;
huffman
parents: 44909
diff changeset
  4052
  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
  4053
    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
  4054
  thus ?thesis using assms(1)
9583f2b56f85 add lemmas within_empty and tendsto_bot;
huffman
parents: 44909
diff changeset
  4055
    using continuous_within_compose[of x UNIV f g] by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4056
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4057
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4058
lemma continuous_on_compose:
36440
89a70297564d simplify definition of continuous_on; generalize some lemmas
huffman
parents: 36439
diff changeset
  4059
  "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
  4060
  unfolding continuous_on_topological by simp metis
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4061
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4062
lemma uniformly_continuous_on_compose:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4063
  assumes "uniformly_continuous_on s f"  "uniformly_continuous_on (f ` s) g"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4064
  shows "uniformly_continuous_on s (g o f)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4065
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4066
  { fix e::real assume "e>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4067
    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
  4068
    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
  4069
    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
  4070
  thus ?thesis using assms unfolding uniformly_continuous_on_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4071
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4072
44647
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4073
lemmas continuous_on_intros = continuous_on_id continuous_on_const
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4074
  continuous_on_compose continuous_on_norm continuous_on_infnorm
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4075
  continuous_on_add continuous_on_minus continuous_on_diff
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4076
  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
  4077
  continuous_on_inner
44648
897f32a827f2 simplify some proofs about uniform continuity, and add some new ones;
huffman
parents: 44647
diff changeset
  4078
  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
  4079
  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
  4080
  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
  4081
  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
  4082
  uniformly_continuous_on_cmul
44647
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4083
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4084
text{* Continuity in terms of open preimages. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4085
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4086
lemma continuous_at_open:
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4087
  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
  4088
unfolding continuous_within_topological [of x UNIV f, unfolded within_UNIV]
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4089
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
  4090
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4091
lemma continuous_on_open:
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4092
  shows "continuous_on s f \<longleftrightarrow>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4093
        (\<forall>t. openin (subtopology euclidean (f ` s)) t
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4094
            --> 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
  4095
proof (safe)
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4096
  fix t :: "'b set"
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4097
  assume 1: "continuous_on s f"
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4098
  assume 2: "openin (subtopology euclidean (f ` s)) t"
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4099
  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
  4100
    unfolding openin_open by auto
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4101
  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
  4102
  have "open U" unfolding U_def by (simp add: open_Union)
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4103
  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
  4104
  proof (intro ballI iffI)
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4105
    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
  4106
      unfolding U_def t by auto
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4107
  next
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4108
    fix x assume "x \<in> s" and "f x \<in> t"
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4109
    hence "x \<in> s" and "f x \<in> B"
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4110
      unfolding t by auto
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4111
    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
  4112
      unfolding t continuous_on_topological by metis
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4113
    then show "x \<in> U"
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4114
      unfolding U_def by auto
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4115
  qed
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4116
  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
  4117
  then show "openin (subtopology euclidean s) {x \<in> s. f x \<in> t}"
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4118
    unfolding openin_open by fast
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4119
next
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4120
  assume "?rhs" show "continuous_on s f"
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4121
  unfolding continuous_on_topological
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4122
  proof (clarify)
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4123
    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
  4124
    have "openin (subtopology euclidean (f ` s)) (f ` s \<inter> B)"
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4125
      unfolding openin_open using `open B` by auto
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4126
    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
  4127
      using `?rhs` by fast
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4128
    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
  4129
      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
  4130
  qed
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4131
qed
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4132
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4133
text {* Similarly in terms of closed sets. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4134
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4135
lemma continuous_on_closed:
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  4136
  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
  4137
proof
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4138
  assume ?lhs
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4139
  { fix t
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4140
    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
  4141
    have **:"f ` s - (f ` s - (f ` s - t)) = f ` s - t" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4142
    assume as:"closedin (subtopology euclidean (f ` s)) t"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4143
    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
  4144
    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
  4145
      unfolding openin_closedin_eq topspace_euclidean_subtopology unfolding * by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4146
  thus ?rhs by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4147
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4148
  assume ?rhs
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4149
  { fix t
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4150
    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
  4151
    assume as:"openin (subtopology euclidean (f ` s)) t"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4152
    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
  4153
      unfolding openin_closedin_eq topspace_euclidean_subtopology *[THEN sym] closedin_subtopology by auto }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4154
  thus ?lhs unfolding continuous_on_open by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4155
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4156
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  4157
text {* Half-global and completely global cases. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4158
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4159
lemma continuous_open_in_preimage:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4160
  assumes "continuous_on s f"  "open t"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4161
  shows "openin (subtopology euclidean s) {x \<in> s. f x \<in> t}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4162
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4163
  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
  4164
  have "openin (subtopology euclidean (f ` s)) (t \<inter> f ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4165
    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
  4166
  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
  4167
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4168
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4169
lemma continuous_closed_in_preimage:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4170
  assumes "continuous_on s f"  "closed t"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4171
  shows "closedin (subtopology euclidean s) {x \<in> s. f x \<in> t}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4172
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4173
  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
  4174
  have "closedin (subtopology euclidean (f ` s)) (t \<inter> f ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4175
    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
  4176
  thus ?thesis
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4177
    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
  4178
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4179
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4180
lemma continuous_open_preimage:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4181
  assumes "continuous_on s f" "open s" "open t"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4182
  shows "open {x \<in> s. f x \<in> t}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4183
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4184
  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
  4185
    using continuous_open_in_preimage[OF assms(1,3)] unfolding openin_open by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4186
  thus ?thesis using open_Int[of s T, OF assms(2)] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4187
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4188
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4189
lemma continuous_closed_preimage:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4190
  assumes "continuous_on s f" "closed s" "closed t"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4191
  shows "closed {x \<in> s. f x \<in> t}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4192
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4193
  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
  4194
    using continuous_closed_in_preimage[OF assms(1,3)] unfolding closedin_closed by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4195
  thus ?thesis using closed_Int[of s T, OF assms(2)] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4196
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4197
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4198
lemma continuous_open_preimage_univ:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4199
  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
  4200
  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
  4201
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4202
lemma continuous_closed_preimage_univ:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4203
  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
  4204
  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
  4205
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4206
lemma continuous_open_vimage:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4207
  shows "\<forall>x. continuous (at x) f \<Longrightarrow> open s \<Longrightarrow> open (f -` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4208
  unfolding vimage_def by (rule continuous_open_preimage_univ)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4209
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4210
lemma continuous_closed_vimage:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4211
  shows "\<forall>x. continuous (at x) f \<Longrightarrow> closed s \<Longrightarrow> closed (f -` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4212
  unfolding vimage_def by (rule continuous_closed_preimage_univ)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4213
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4214
lemma interior_image_subset:
35172
579dd5570f96 Added integration to Multivariate-Analysis (upto FTC)
himmelma
parents: 35028
diff changeset
  4215
  assumes "\<forall>x. continuous (at x) f" "inj f"
579dd5570f96 Added integration to Multivariate-Analysis (upto FTC)
himmelma
parents: 35028
diff changeset
  4216
  shows "interior (f ` s) \<subseteq> f ` (interior s)"
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  4217
proof
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  4218
  fix x assume "x \<in> interior (f ` s)"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  4219
  then obtain T where as: "open T" "x \<in> T" "T \<subseteq> f ` s" ..
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  4220
  hence "x \<in> f ` s" by auto
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  4221
  then obtain y where y: "y \<in> s" "x = f y" by auto
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  4222
  have "open (vimage f T)"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  4223
    using assms(1) `open T` by (rule continuous_open_vimage)
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  4224
  moreover have "y \<in> vimage f T"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  4225
    using `x = f y` `x \<in> T` by simp
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  4226
  moreover have "vimage f T \<subseteq> s"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  4227
    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
  4228
  ultimately have "y \<in> interior s" ..
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  4229
  with `x = f y` show "x \<in> f ` interior s" ..
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  4230
qed
35172
579dd5570f96 Added integration to Multivariate-Analysis (upto FTC)
himmelma
parents: 35028
diff changeset
  4231
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  4232
text {* Equality of continuous functions on closure and related results. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4233
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4234
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
  4235
  fixes f :: "_ \<Rightarrow> 'b::t1_space"
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  4236
  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
  4237
  using continuous_closed_in_preimage[of s f "{a}"] by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4238
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4239
lemma continuous_closed_preimage_constant:
36668
941ba2da372e simplify definition of t1_space; generalize lemma closed_sing and related lemmas
huffman
parents: 36667
diff changeset
  4240
  fixes f :: "_ \<Rightarrow> 'b::t1_space"
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  4241
  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
  4242
  using continuous_closed_preimage[of s f "{a}"] by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4243
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4244
lemma continuous_constant_on_closure:
36668
941ba2da372e simplify definition of t1_space; generalize lemma closed_sing and related lemmas
huffman
parents: 36667
diff changeset
  4245
  fixes f :: "_ \<Rightarrow> 'b::t1_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4246
  assumes "continuous_on (closure s) f"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4247
          "\<forall>x \<in> s. f x = a"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4248
  shows "\<forall>x \<in> (closure s). f x = a"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4249
    using continuous_closed_preimage_constant[of "closure s" f a]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4250
    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
  4251
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4252
lemma image_closure_subset:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4253
  assumes "continuous_on (closure s) f"  "closed t"  "(f ` s) \<subseteq> t"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4254
  shows "f ` (closure s) \<subseteq> t"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4255
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4256
  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
  4257
  moreover have "closed {x \<in> closure s. f x \<in> t}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4258
    using continuous_closed_preimage[OF assms(1)] and assms(2) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4259
  ultimately have "closure s = {x \<in> closure s . f x \<in> t}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4260
    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
  4261
  thus ?thesis by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4262
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4263
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4264
lemma continuous_on_closure_norm_le:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4265
  fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4266
  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
  4267
  shows "norm(f x) \<le> b"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4268
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4269
  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
  4270
  show ?thesis
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4271
    using image_closure_subset[OF assms(1) closed_cball[of 0 b] *] assms(3)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4272
    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
  4273
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4274
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  4275
text {* Making a continuous function avoid some value in a neighbourhood. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4276
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4277
lemma continuous_within_avoid:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4278
  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" (* FIXME: generalize *)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4279
  assumes "continuous (at x within s) f"  "x \<in> s"  "f x \<noteq> a"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4280
  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
  4281
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4282
  obtain d where "d>0" and d:"\<forall>xa\<in>s. 0 < dist xa x \<and> dist xa x < d \<longrightarrow> dist (f xa) (f x) < dist (f x) a"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4283
    using assms(1)[unfolded continuous_within Lim_within, THEN spec[where x="dist (f x) a"]] assms(3)[unfolded dist_nz] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4284
  { fix y assume " y\<in>s"  "dist x y < d"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4285
    hence "f y \<noteq> a" using d[THEN bspec[where x=y]] assms(3)[unfolded dist_nz]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4286
      apply auto unfolding dist_nz[THEN sym] by (auto simp add: dist_commute) }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4287
  thus ?thesis using `d>0` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4288
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4289
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4290
lemma continuous_at_avoid:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4291
  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" (* FIXME: generalize *)
45031
9583f2b56f85 add lemmas within_empty and tendsto_bot;
huffman
parents: 44909
diff changeset
  4292
  assumes "continuous (at x) f" and "f x \<noteq> a"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4293
  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
  4294
  using assms continuous_within_avoid[of x UNIV f a] by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4295
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4296
lemma continuous_on_avoid:
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  4297
  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" (* TODO: generalize *)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4298
  assumes "continuous_on s f"  "x \<in> s"  "f x \<noteq> a"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4299
  shows "\<exists>e>0. \<forall>y \<in> s. dist x y < e \<longrightarrow> f y \<noteq> a"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4300
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(2,3) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4301
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4302
lemma continuous_on_open_avoid:
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  4303
  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" (* TODO: generalize *)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4304
  assumes "continuous_on s f"  "open s"  "x \<in> s"  "f x \<noteq> a"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4305
  shows "\<exists>e>0. \<forall>y. dist x y < e \<longrightarrow> f y \<noteq> a"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4306
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(3,4) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4307
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  4308
text {* Proving a function is constant by proving open-ness of level set. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4309
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4310
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
  4311
  fixes f :: "_ \<Rightarrow> 'b::t1_space"
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  4312
  shows "connected s \<Longrightarrow> continuous_on s f \<Longrightarrow>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4313
        openin (subtopology euclidean s) {x \<in> s. f x = a}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4314
        ==> (\<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
  4315
unfolding connected_clopen using continuous_closed_in_preimage_constant by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4316
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4317
lemma continuous_levelset_open_in:
36668
941ba2da372e simplify definition of t1_space; generalize lemma closed_sing and related lemmas
huffman
parents: 36667
diff changeset
  4318
  fixes f :: "_ \<Rightarrow> 'b::t1_space"
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  4319
  shows "connected s \<Longrightarrow> continuous_on s f \<Longrightarrow>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4320
        openin (subtopology euclidean s) {x \<in> s. f x = a} \<Longrightarrow>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4321
        (\<exists>x \<in> s. f x = a)  ==> (\<forall>x \<in> s. f x = a)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4322
using continuous_levelset_open_in_cases[of s f ]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4323
by meson
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4324
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4325
lemma continuous_levelset_open:
36668
941ba2da372e simplify definition of t1_space; generalize lemma closed_sing and related lemmas
huffman
parents: 36667
diff changeset
  4326
  fixes f :: "_ \<Rightarrow> 'b::t1_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4327
  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
  4328
  shows "\<forall>x \<in> s. f x = a"
36362
06475a1547cb fix lots of looping simp calls and other warnings
huffman
parents: 36360
diff changeset
  4329
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
  4330
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  4331
text {* Some arithmetical combinations (more to prove). *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4332
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4333
lemma open_scaling[intro]:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4334
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4335
  assumes "c \<noteq> 0"  "open s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4336
  shows "open((\<lambda>x. c *\<^sub>R x) ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4337
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4338
  { fix x assume "x \<in> s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4339
    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
  4340
    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
  4341
    moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4342
    { fix y assume "dist y (c *\<^sub>R x) < e * \<bar>c\<bar>"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4343
      hence "norm ((1 / c) *\<^sub>R y - x) < e" unfolding dist_norm
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4344
        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
  4345
          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
  4346
      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
  4347
    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
  4348
  thus ?thesis unfolding open_dist by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4349
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4350
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4351
lemma minus_image_eq_vimage:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4352
  fixes A :: "'a::ab_group_add set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4353
  shows "(\<lambda>x. - x) ` A = (\<lambda>x. - x) -` A"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4354
  by (auto intro!: image_eqI [where f="\<lambda>x. - x"])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4355
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4356
lemma open_negations:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4357
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4358
  shows "open s ==> open ((\<lambda> x. -x) ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4359
  unfolding scaleR_minus1_left [symmetric]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4360
  by (rule open_scaling, auto)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4361
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4362
lemma open_translation:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4363
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4364
  assumes "open s"  shows "open((\<lambda>x. a + x) ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4365
proof-
44647
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4366
  { fix x have "continuous (at x) (\<lambda>x. x - a)"
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4367
      by (intro continuous_diff continuous_at_id continuous_const) }
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4368
  moreover have "{x. x - a \<in> s} = op + a ` s" by force
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4369
  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
  4370
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4371
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4372
lemma open_affinity:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4373
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4374
  assumes "open s"  "c \<noteq> 0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4375
  shows "open ((\<lambda>x. a + c *\<^sub>R x) ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4376
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4377
  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
  4378
  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
  4379
  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
  4380
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4381
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4382
lemma interior_translation:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4383
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4384
  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
  4385
proof (rule set_eqI, rule)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4386
  fix x assume "x \<in> interior (op + a ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4387
  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
  4388
  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
  4389
  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
  4390
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4391
  fix x assume "x \<in> op + a ` interior s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4392
  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
  4393
  { fix z have *:"a + y - z = y + a - z" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4394
    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
  4395
    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
  4396
    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
  4397
  hence "ball x e \<subseteq> op + a ` s" unfolding subset_eq by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4398
  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
  4399
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4400
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4401
text {* Topological properties of linear functions. *}
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4402
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4403
lemma linear_lim_0:
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4404
  assumes "bounded_linear f" shows "(f ---> 0) (at (0))"
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4405
proof-
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4406
  interpret f: bounded_linear f by fact
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4407
  have "(f ---> f 0) (at 0)"
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4408
    using tendsto_ident_at by (rule f.tendsto)
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4409
  thus ?thesis unfolding f.zero .
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4410
qed
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4411
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4412
lemma linear_continuous_at:
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4413
  assumes "bounded_linear f"  shows "continuous (at a) f"
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4414
  unfolding continuous_at using assms
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4415
  apply (rule bounded_linear.tendsto)
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4416
  apply (rule tendsto_ident_at)
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4417
  done
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4418
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4419
lemma linear_continuous_within:
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4420
  shows "bounded_linear f ==> continuous (at x within s) f"
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4421
  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
  4422
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4423
lemma linear_continuous_on:
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4424
  shows "bounded_linear f ==> continuous_on s f"
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4425
  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
  4426
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  4427
text {* Also bilinear functions, in composition form. *}
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4428
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4429
lemma bilinear_continuous_at_compose:
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4430
  shows "continuous (at x) f \<Longrightarrow> continuous (at x) g \<Longrightarrow> bounded_bilinear h
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4431
        ==> continuous (at x) (\<lambda>x. h (f x) (g x))"
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4432
  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
  4433
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4434
lemma bilinear_continuous_within_compose:
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4435
  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
  4436
        ==> continuous (at x within s) (\<lambda>x. h (f x) (g x))"
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4437
  unfolding continuous_within using Lim_bilinear[of f "f x"] by auto
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4438
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4439
lemma bilinear_continuous_on_compose:
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4440
  shows "continuous_on s f \<Longrightarrow> continuous_on s g \<Longrightarrow> bounded_bilinear h
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4441
             ==> continuous_on s (\<lambda>x. h (f x) (g x))"
36441
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4442
  unfolding continuous_on_def
1d7704c29af3 generalized many lemmas about continuity
huffman
parents: 36440
diff changeset
  4443
  by (fast elim: bounded_bilinear.tendsto)
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4444
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  4445
text {* Preservation of compactness and connectedness under continuous function. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4446
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4447
lemma compact_continuous_image:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4448
  assumes "continuous_on s f"  "compact s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4449
  shows "compact(f ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4450
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4451
  { fix x assume x:"\<forall>n::nat. x n \<in> f ` s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4452
    then obtain y where y:"\<forall>n. y n \<in> s \<and> x n = f (y n)" unfolding image_iff Bex_def using choice[of "\<lambda>n xa. xa \<in> s \<and> x n = f xa"] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4453
    then obtain l r where "l\<in>s" and r:"subseq r" and lr:"((y \<circ> r) ---> l) sequentially" using assms(2)[unfolded compact_def, THEN spec[where x=y]] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4454
    { fix e::real assume "e>0"
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  4455
      then obtain d where "d>0" and d:"\<forall>x'\<in>s. dist x' l < d \<longrightarrow> dist (f x') (f l) < e" using assms(1)[unfolded continuous_on_iff, THEN bspec[where x=l], OF `l\<in>s`] by auto
44907
93943da0a010 remove redundant lemma Lim_sequentially in favor of lemma LIMSEQ_def
huffman
parents: 44905
diff changeset
  4456
      then obtain N::nat where N:"\<forall>n\<ge>N. dist ((y \<circ> r) n) l < d" using lr[unfolded LIMSEQ_def, THEN spec[where x=d]] by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4457
      { fix n::nat assume "n\<ge>N" hence "dist ((x \<circ> r) n) (f l) < e" using N[THEN spec[where x=n]] d[THEN bspec[where x="y (r n)"]] y[THEN spec[where x="r n"]] by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4458
      hence "\<exists>N. \<forall>n\<ge>N. dist ((x \<circ> r) n) (f l) < e" by auto  }
44907
93943da0a010 remove redundant lemma Lim_sequentially in favor of lemma LIMSEQ_def
huffman
parents: 44905
diff changeset
  4459
    hence "\<exists>l\<in>f ` s. \<exists>r. subseq r \<and> ((x \<circ> r) ---> l) sequentially" unfolding LIMSEQ_def using r lr `l\<in>s` by auto  }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4460
  thus ?thesis unfolding compact_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4461
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4462
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4463
lemma connected_continuous_image:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4464
  assumes "continuous_on s f"  "connected s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4465
  shows "connected(f ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4466
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4467
  { 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
  4468
    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
  4469
      using assms(1)[unfolded continuous_on_open, THEN spec[where x=T]]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4470
      using assms(1)[unfolded continuous_on_closed, THEN spec[where x=T]]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4471
      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
  4472
    hence False using as(1,2)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4473
      using as(4)[unfolded closedin_def topspace_euclidean_subtopology] by auto }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4474
  thus ?thesis unfolding connected_clopen by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4475
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4476
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  4477
text {* Continuity implies uniform continuity on a compact domain. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4478
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4479
lemma compact_uniformly_continuous:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4480
  assumes "continuous_on s f"  "compact s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4481
  shows "uniformly_continuous_on s f"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4482
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4483
    { fix x assume x:"x\<in>s"
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  4484
      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
  4485
      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
  4486
    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
  4487
    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
  4488
      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
  4489
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4490
  { fix e::real assume "e>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4491
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4492
    { 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
  4493
    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
  4494
    moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4495
    { fix b assume "b\<in>{ball x (d x (e / 2)) |x. x \<in> s}" hence "open b" by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4496
    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" using heine_borel_lemma[OF assms(2), of "{ball x (d x (e / 2)) | x. x\<in>s }"] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4497
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4498
    { 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
  4499
      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
  4500
      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
  4501
      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
  4502
        by (auto  simp add: dist_commute)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4503
      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
  4504
        by (auto simp add: dist_commute)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4505
      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
  4506
        by (auto  simp add: dist_commute)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4507
      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
  4508
        by (auto simp add: dist_commute)  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4509
    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
  4510
  thus ?thesis unfolding uniformly_continuous_on_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4511
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4512
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4513
text{* Continuity of inverse function on compact domain. *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4514
44647
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  4515
lemma continuous_on_inv:
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4516
  fixes f :: "'a::heine_borel \<Rightarrow> 'b::heine_borel"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4517
    (* TODO: can this be generalized more? *)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4518
  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
  4519
  shows "continuous_on (f ` s) g"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4520
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4521
  have *:"g ` f ` s = s" using assms(3) by (auto simp add: image_iff)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4522
  { fix t assume t:"closedin (subtopology euclidean (g ` f ` s)) t"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4523
    then obtain T where T: "closed T" "t = s \<inter> T" unfolding closedin_closed unfolding * by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4524
    have "continuous_on (s \<inter> T) f" using continuous_on_subset[OF assms(1), of "s \<inter> t"]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4525
      unfolding T(2) and Int_left_absorb by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4526
    moreover have "compact (s \<inter> T)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4527
      using assms(2) unfolding compact_eq_bounded_closed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4528
      using bounded_subset[of s "s \<inter> T"] and T(1) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4529
    ultimately have "closed (f ` t)" using T(1) unfolding T(2)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4530
      using compact_continuous_image [of "s \<inter> T" f] unfolding compact_eq_bounded_closed by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4531
    moreover have "{x \<in> f ` s. g x \<in> t} = f ` s \<inter> f ` t" using assms(3) unfolding T(2) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4532
    ultimately have "closedin (subtopology euclidean (f ` s)) {x \<in> f ` s. g x \<in> t}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4533
      unfolding closedin_closed by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4534
  thus ?thesis unfolding continuous_on_closed by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4535
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4536
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4537
text {* A uniformly convergent limit of continuous functions is continuous. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4538
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4539
lemma continuous_uniform_limit:
44212
4d10e7f342b1 generalize lemma continuous_uniform_limit to class metric_space
huffman
parents: 44211
diff changeset
  4540
  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
  4541
  assumes "\<not> trivial_limit F"
4d10e7f342b1 generalize lemma continuous_uniform_limit to class metric_space
huffman
parents: 44211
diff changeset
  4542
  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
  4543
  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
  4544
  shows "continuous_on s g"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4545
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4546
  { 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
  4547
    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
  4548
      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
  4549
    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
  4550
    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
  4551
      using assms(1) by blast
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4552
    have "e / 3 > 0" using `e>0` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4553
    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
  4554
      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
  4555
    { 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
  4556
      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
  4557
        by (rule d [rule_format])
4d10e7f342b1 generalize lemma continuous_uniform_limit to class metric_space
huffman
parents: 44211
diff changeset
  4558
      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
  4559
        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
  4560
        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
  4561
        by auto
4d10e7f342b1 generalize lemma continuous_uniform_limit to class metric_space
huffman
parents: 44211
diff changeset
  4562
      hence "dist (g y) (g x) < e"
4d10e7f342b1 generalize lemma continuous_uniform_limit to class metric_space
huffman
parents: 44211
diff changeset
  4563
        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
  4564
        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
  4565
        by auto }
4d10e7f342b1 generalize lemma continuous_uniform_limit to class metric_space
huffman
parents: 44211
diff changeset
  4566
    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
  4567
      using `d>0` by auto }
36359
e5c785c166b2 generalize type of continuous_on
huffman
parents: 36358
diff changeset
  4568
  thus ?thesis unfolding continuous_on_iff by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4569
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4570
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  4571
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  4572
subsection {* Topological stuff lifted from and dropped to R *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4573
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4574
lemma open_real:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4575
  fixes s :: "real set" shows
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4576
 "open s \<longleftrightarrow>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4577
        (\<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
  4578
  unfolding open_dist dist_norm by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4579
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4580
lemma islimpt_approachable_real:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4581
  fixes s :: "real set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4582
  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
  4583
  unfolding islimpt_approachable dist_norm by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4584
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4585
lemma closed_real:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4586
  fixes s :: "real set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4587
  shows "closed s \<longleftrightarrow>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4588
        (\<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
  4589
            --> x \<in> s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4590
  unfolding closed_limpt islimpt_approachable dist_norm by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4591
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4592
lemma continuous_at_real_range:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4593
  fixes f :: "'a::real_normed_vector \<Rightarrow> real"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4594
  shows "continuous (at x) f \<longleftrightarrow> (\<forall>e>0. \<exists>d>0.
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4595
        \<forall>x'. norm(x' - x) < d --> abs(f x' - f x) < e)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4596
  unfolding continuous_at unfolding Lim_at
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4597
  unfolding dist_nz[THEN sym] unfolding dist_norm apply auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4598
  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
  4599
  apply(erule_tac x=e in allE) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4600
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4601
lemma continuous_on_real_range:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4602
  fixes f :: "'a::real_normed_vector \<Rightarrow> real"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4603
  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
  4604
  unfolding continuous_on_iff dist_norm by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4605
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  4606
text {* Hence some handy theorems on distance, diameter etc. of/from a set. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4607
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4608
lemma compact_attains_sup:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4609
  fixes s :: "real set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4610
  assumes "compact s"  "s \<noteq> {}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4611
  shows "\<exists>x \<in> s. \<forall>y \<in> s. y \<le> x"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4612
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4613
  from assms(1) have a:"bounded s" "closed s" unfolding compact_eq_bounded_closed by auto
33270
paulson
parents: 33175
diff changeset
  4614
  { 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
  4615
    have "isLub UNIV s (Sup s)" using Sup[OF assms(2)] unfolding setle_def using as(1) by auto
paulson
parents: 33175
diff changeset
  4616
    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
  4617
    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
  4618
  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
  4619
    apply(rule_tac x="Sup s" in bexI) by auto
paulson
parents: 33175
diff changeset
  4620
qed
paulson
parents: 33175
diff changeset
  4621
paulson
parents: 33175
diff changeset
  4622
lemma Inf:
paulson
parents: 33175
diff changeset
  4623
  fixes S :: "real set"
paulson
parents: 33175
diff changeset
  4624
  shows "S \<noteq> {} ==> (\<exists>b. b <=* S) ==> isGlb UNIV S (Inf S)"
paulson
parents: 33175
diff changeset
  4625
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
  4626
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4627
lemma compact_attains_inf:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4628
  fixes s :: "real set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4629
  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
  4630
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4631
  from assms(1) have a:"bounded s" "closed s" unfolding compact_eq_bounded_closed by auto
33270
paulson
parents: 33175
diff changeset
  4632
  { fix e::real assume as: "\<forall>x\<in>s. x \<ge> Inf s"  "Inf s \<notin> s"  "0 < e"
paulson
parents: 33175
diff changeset
  4633
      "\<forall>x'\<in>s. x' = Inf s \<or> \<not> abs (x' - Inf s) < e"
paulson
parents: 33175
diff changeset
  4634
    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
  4635
    moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4636
    { fix x assume "x \<in> s"
33270
paulson
parents: 33175
diff changeset
  4637
      hence *:"abs (x - Inf s) = x - Inf s" using as(1)[THEN bspec[where x=x]] by auto
paulson
parents: 33175
diff changeset
  4638
      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
  4639
    hence "isLb UNIV s (Inf s + e)" unfolding isLb_def and setge_def by auto
paulson
parents: 33175
diff changeset
  4640
    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
  4641
  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
  4642
    apply(rule_tac x="Inf s" in bexI) by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4643
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4644
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4645
lemma continuous_attains_sup:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4646
  fixes f :: "'a::metric_space \<Rightarrow> real"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4647
  shows "compact s \<Longrightarrow> s \<noteq> {} \<Longrightarrow> continuous_on s f
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4648
        ==> (\<exists>x \<in> s. \<forall>y \<in> s.  f y \<le> f x)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4649
  using compact_attains_sup[of "f ` s"]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4650
  using compact_continuous_image[of s f] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4651
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4652
lemma continuous_attains_inf:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4653
  fixes f :: "'a::metric_space \<Rightarrow> real"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4654
  shows "compact s \<Longrightarrow> s \<noteq> {} \<Longrightarrow> continuous_on s f
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4655
        \<Longrightarrow> (\<exists>x \<in> s. \<forall>y \<in> s. f x \<le> f y)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4656
  using compact_attains_inf[of "f ` s"]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4657
  using compact_continuous_image[of s f] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4658
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4659
lemma distance_attains_sup:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4660
  assumes "compact s" "s \<noteq> {}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4661
  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
  4662
proof (rule continuous_attains_sup [OF assms])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4663
  { fix x assume "x\<in>s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4664
    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
  4665
      by (intro tendsto_dist tendsto_const Lim_at_within tendsto_ident_at)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4666
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4667
  thus "continuous_on s (dist a)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4668
    unfolding continuous_on ..
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4669
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4670
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  4671
text {* For \emph{minimal} distance, we only need closure, not compactness. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4672
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4673
lemma distance_attains_inf:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4674
  fixes a :: "'a::heine_borel"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4675
  assumes "closed s"  "s \<noteq> {}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4676
  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
  4677
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4678
  from assms(2) obtain b where "b\<in>s" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4679
  let ?B = "cball a (dist b a) \<inter> s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4680
  have "b \<in> ?B" using `b\<in>s` by (simp add: dist_commute)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4681
  hence "?B \<noteq> {}" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4682
  moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4683
  { fix x assume "x\<in>?B"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4684
    fix e::real assume "e>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4685
    { fix x' assume "x'\<in>?B" and as:"dist x' x < e"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4686
      from as have "\<bar>dist a x' - dist a x\<bar> < e"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4687
        unfolding abs_less_iff minus_diff_eq
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4688
        using dist_triangle2 [of a x' x]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4689
        using dist_triangle [of a x x']
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4690
        by arith
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4691
    }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4692
    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
  4693
      using `e>0` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4694
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4695
  hence "continuous_on (cball a (dist b a) \<inter> s) (dist a)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4696
    unfolding continuous_on Lim_within dist_norm real_norm_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4697
    by fast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4698
  moreover have "compact ?B"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4699
    using compact_cball[of a "dist b a"]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4700
    unfolding compact_eq_bounded_closed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4701
    using bounded_Int and closed_Int and assms(1) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4702
  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
  4703
    using continuous_attains_inf[of ?B "dist a"] by fastforce
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 44668
diff changeset
  4704
  thus ?thesis by fastforce
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4705
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4706
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  4707
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  4708
subsection {* Pasted sets *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4709
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4710
lemma bounded_Times:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4711
  assumes "bounded s" "bounded t" shows "bounded (s \<times> t)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4712
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4713
  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
  4714
    using assms [unfolded bounded_def] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4715
  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
  4716
    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
  4717
  thus ?thesis unfolding bounded_any_center [where a="(x, y)"] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4718
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4719
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4720
lemma 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
  4721
by (induct x) simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4722
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4723
lemma compact_Times: "compact s \<Longrightarrow> compact t \<Longrightarrow> compact (s \<times> t)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4724
unfolding compact_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4725
apply clarify
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4726
apply (drule_tac x="fst \<circ> f" in spec)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4727
apply (drule mp, simp add: mem_Times_iff)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4728
apply (clarify, rename_tac l1 r1)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4729
apply (drule_tac x="snd \<circ> f \<circ> r1" in spec)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4730
apply (drule mp, simp add: mem_Times_iff)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4731
apply (clarify, rename_tac l2 r2)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4732
apply (rule_tac x="(l1, l2)" in rev_bexI, simp)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4733
apply (rule_tac x="r1 \<circ> r2" in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4734
apply (rule conjI, simp add: subseq_def)
48125
602dc0215954 tuned proofs -- prefer direct "rotated" instead of old-style COMP;
wenzelm
parents: 48048
diff changeset
  4735
apply (drule_tac r=r2 in lim_subseq [rotated], assumption)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4736
apply (drule (1) tendsto_Pair) back
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4737
apply (simp add: o_def)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4738
done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4739
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  4740
text{* Hence some useful properties follow quite easily. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4741
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4742
lemma compact_scaling:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4743
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4744
  assumes "compact s"  shows "compact ((\<lambda>x. c *\<^sub>R x) ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4745
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4746
  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
  4747
  have *:"bounded_linear ?f" by (rule bounded_linear_scaleR_right)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4748
  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
  4749
    using linear_continuous_at[OF *] assms 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
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4752
lemma compact_negations:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4753
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4754
  assumes "compact s"  shows "compact ((\<lambda>x. -x) ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4755
  using compact_scaling [OF assms, of "- 1"] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4756
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4757
lemma compact_sums:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4758
  fixes s t :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4759
  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
  4760
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4761
  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
  4762
    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
  4763
  have "continuous_on (s \<times> t) (\<lambda>z. fst z + snd z)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4764
    unfolding continuous_on by (rule ballI) (intro tendsto_intros)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4765
  thus ?thesis unfolding * using compact_continuous_image compact_Times [OF assms] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4766
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4767
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4768
lemma compact_differences:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4769
  fixes s t :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4770
  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
  4771
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4772
  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
  4773
    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
  4774
  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
  4775
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4776
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4777
lemma compact_translation:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4778
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4779
  assumes "compact s"  shows "compact ((\<lambda>x. a + x) ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4780
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4781
  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
  4782
  thus ?thesis using compact_sums[OF assms compact_sing[of a]] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4783
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4784
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4785
lemma compact_affinity:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4786
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4787
  assumes "compact s"  shows "compact ((\<lambda>x. a + c *\<^sub>R x) ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4788
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4789
  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
  4790
  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
  4791
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4792
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  4793
text {* Hence we get the following. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4794
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4795
lemma compact_sup_maxdistance:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4796
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4797
  assumes "compact s"  "s \<noteq> {}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4798
  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
  4799
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4800
  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
  4801
  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
  4802
    using compact_differences[OF assms(1) assms(1)]
36362
06475a1547cb fix lots of looping simp calls and other warnings
huffman
parents: 36360
diff changeset
  4803
    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
  4804
  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
  4805
  thus ?thesis using x(2)[unfolded `x = a - b`] by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4806
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4807
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  4808
text {* We can state this in terms of diameter of a set. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4809
33270
paulson
parents: 33175
diff changeset
  4810
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
  4811
  (* TODO: generalize to class metric_space *)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4812
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4813
lemma diameter_bounded:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4814
  assumes "bounded s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4815
  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
  4816
        "\<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
  4817
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4818
  let ?D = "{norm (x - y) |x y. x \<in> s \<and> y \<in> s}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4819
  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
  4820
  { 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
  4821
    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
  4822
  note * = this
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4823
  { 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
  4824
    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
  4825
      by simp (blast del: Sup_upper intro!: * Sup_upper) }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4826
  moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4827
  { fix d::real assume "d>0" "d < diameter s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4828
    hence "s\<noteq>{}" unfolding diameter_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4829
    have "\<exists>d' \<in> ?D. d' > d"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4830
    proof(rule ccontr)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4831
      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
  4832
      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
  4833
      thus False using `d < diameter s` `s\<noteq>{}` 
51eb2ffa2189 Tidied up some very ugly proofs
paulson
parents: 33270
diff changeset
  4834
        apply (auto simp add: diameter_def) 
51eb2ffa2189 Tidied up some very ugly proofs
paulson
parents: 33270
diff changeset
  4835
        apply (drule Sup_real_iff [THEN [2] rev_iffD2])
51eb2ffa2189 Tidied up some very ugly proofs
paulson
parents: 33270
diff changeset
  4836
        apply (auto, force) 
51eb2ffa2189 Tidied up some very ugly proofs
paulson
parents: 33270
diff changeset
  4837
        done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4838
    qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4839
    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
  4840
  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
  4841
        "\<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
  4842
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4843
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4844
lemma diameter_bounded_bound:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4845
 "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
  4846
  using diameter_bounded by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4847
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4848
lemma diameter_compact_attained:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4849
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4850
  assumes "compact s"  "s \<noteq> {}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4851
  shows "\<exists>x\<in>s. \<exists>y\<in>s. (norm(x - y) = diameter s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4852
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4853
  have b:"bounded s" using assms(1) by (rule compact_imp_bounded)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4854
  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
  4855
  hence "diameter s \<le> norm (x - y)"
06475a1547cb fix lots of looping simp calls and other warnings
huffman
parents: 36360
diff changeset
  4856
    unfolding diameter_def by clarsimp (rule Sup_least, fast+)
33324
51eb2ffa2189 Tidied up some very ugly proofs
paulson
parents: 33270
diff changeset
  4857
  thus ?thesis
36362
06475a1547cb fix lots of looping simp calls and other warnings
huffman
parents: 36360
diff changeset
  4858
    by (metis b diameter_bounded_bound order_antisym xys)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4859
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4860
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  4861
text {* Related results with closure as the conclusion. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4862
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4863
lemma closed_scaling:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4864
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4865
  assumes "closed s" shows "closed ((\<lambda>x. c *\<^sub>R x) ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4866
proof(cases "s={}")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4867
  case True thus ?thesis by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4868
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4869
  case False
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4870
  show ?thesis
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4871
  proof(cases "c=0")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4872
    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
  4873
    case True thus ?thesis apply auto unfolding * by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4874
  next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4875
    case False
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4876
    { 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
  4877
      { fix n::nat have "scaleR (1 / c) (x n) \<in> s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4878
          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
  4879
          using `c\<noteq>0` by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4880
      }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4881
      moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4882
      { fix e::real assume "e>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4883
        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
  4884
        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
  4885
          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
  4886
        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
  4887
          unfolding dist_norm unfolding scaleR_right_diff_distrib[THEN sym]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4888
          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
  4889
      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
  4890
      ultimately have "l \<in> scaleR c ` s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4891
        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
  4892
        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
  4893
    thus ?thesis unfolding closed_sequential_limits by fast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4894
  qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4895
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4896
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4897
lemma closed_negations:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4898
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4899
  assumes "closed s"  shows "closed ((\<lambda>x. -x) ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4900
  using closed_scaling[OF assms, of "- 1"] by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4901
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4902
lemma compact_closed_sums:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4903
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4904
  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
  4905
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4906
  let ?S = "{x + y |x y. x \<in> s \<and> y \<in> t}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4907
  { fix x l assume as:"\<forall>n. x n \<in> ?S"  "(x ---> l) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4908
    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
  4909
      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
  4910
    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
  4911
      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
  4912
    have "((\<lambda>n. snd (f (r n))) ---> l - l') sequentially"
44125
230a8665c919 mark some redundant theorems as legacy
huffman
parents: 44122
diff changeset
  4913
      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
  4914
    hence "l - l' \<in> t"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4915
      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
  4916
      using f(3) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4917
    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
  4918
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4919
  thus ?thesis unfolding closed_sequential_limits by fast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4920
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4921
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4922
lemma closed_compact_sums:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4923
  fixes s t :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4924
  assumes "closed s"  "compact t"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4925
  shows "closed {x + y | x y. x \<in> s \<and> y \<in> t}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4926
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4927
  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
  4928
    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
  4929
  thus ?thesis using compact_closed_sums[OF assms(2,1)] by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4930
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4931
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4932
lemma compact_closed_differences:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4933
  fixes s t :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4934
  assumes "compact s"  "closed t"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4935
  shows "closed {x - y | x y. x \<in> s \<and> y \<in> t}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4936
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4937
  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
  4938
    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
  4939
  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
  4940
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4941
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4942
lemma closed_compact_differences:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4943
  fixes s t :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4944
  assumes "closed s" "compact t"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4945
  shows "closed {x - y | x y. x \<in> s \<and> y \<in> t}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4946
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4947
  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
  4948
    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
  4949
 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
  4950
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4951
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4952
lemma closed_translation:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4953
  fixes a :: "'a::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4954
  assumes "closed s"  shows "closed ((\<lambda>x. a + x) ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4955
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4956
  have "{a + y |y. y \<in> s} = (op + a ` s)" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4957
  thus ?thesis using compact_closed_sums[OF compact_sing[of a] assms] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4958
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4959
34105
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
  4960
lemma translation_Compl:
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
  4961
  fixes a :: "'a::ab_group_add"
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
  4962
  shows "(\<lambda>x. a + x) ` (- t) = - ((\<lambda>x. a + x) ` t)"
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
  4963
  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
  4964
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4965
lemma translation_UNIV:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4966
  fixes a :: "'a::ab_group_add" shows "range (\<lambda>x. a + x) = UNIV"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4967
  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
  4968
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4969
lemma translation_diff:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4970
  fixes a :: "'a::ab_group_add"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4971
  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
  4972
  by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4973
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4974
lemma closure_translation:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4975
  fixes a :: "'a::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4976
  shows "closure ((\<lambda>x. a + x) ` s) = (\<lambda>x. a + x) ` (closure s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4977
proof-
34105
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
  4978
  have *:"op + a ` (- s) = - op + a ` s"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4979
    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
  4980
  show ?thesis unfolding closure_interior translation_Compl
87cbdecaa879 replace 'UNIV - S' with '- S'
huffman
parents: 34104
diff changeset
  4981
    using interior_translation[of a "- s"] unfolding * by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4982
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4983
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4984
lemma frontier_translation:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4985
  fixes a :: "'a::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4986
  shows "frontier((\<lambda>x. a + x) ` s) = (\<lambda>x. a + x) ` (frontier s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4987
  unfolding frontier_def translation_diff interior_translation closure_translation by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4988
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  4989
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  4990
subsection {* Separation between points and sets *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4991
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4992
lemma separate_point_closed:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4993
  fixes s :: "'a::heine_borel set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4994
  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
  4995
proof(cases "s = {}")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4996
  case True
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4997
  thus ?thesis by(auto intro!: exI[where x=1])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4998
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4999
  case False
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5000
  assume "closed s" "a \<notin> s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5001
  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
  5002
  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
  5003
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5004
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5005
lemma separate_compact_closed:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5006
  fixes s t :: "'a::{heine_borel, real_normed_vector} set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5007
    (* TODO: does this generalize to heine_borel? *)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5008
  assumes "compact s" and "closed t" and "s \<inter> t = {}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5009
  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
  5010
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5011
  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
  5012
  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
  5013
    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
  5014
  { fix x y assume "x\<in>s" "y\<in>t"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5015
    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
  5016
    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
  5017
      by (auto  simp add: dist_commute)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5018
    hence "d \<le> dist x y" unfolding dist_norm by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5019
  thus ?thesis using `d>0` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5020
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5021
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5022
lemma separate_closed_compact:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5023
  fixes s t :: "'a::{heine_borel, real_normed_vector} set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5024
  assumes "closed s" and "compact t" and "s \<inter> t = {}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5025
  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
  5026
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5027
  have *:"t \<inter> s = {}" using assms(3) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5028
  show ?thesis using separate_compact_closed[OF assms(2,1) *]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5029
    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
  5030
    by (auto simp add: dist_commute)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5031
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5032
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  5033
36439
a65320184de9 move intervals section heading
huffman
parents: 36438
diff changeset
  5034
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
  5035
  
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5036
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
  5037
  "{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
  5038
  "{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
  5039
  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
  5040
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5041
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
  5042
  "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
  5043
  "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
  5044
  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
  5045
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5046
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
  5047
 "({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
  5048
 "({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
  5049
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
  5050
  { 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
  5051
    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
  5052
    hence "a\<bullet>i < b\<bullet>i" by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5053
    hence False using as by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5054
  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
  5055
  { 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
  5056
    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
  5057
    { 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
  5058
      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
  5059
      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
  5060
        by (auto simp: inner_add_left) }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5061
    hence "{a <..< b} \<noteq> {}" using mem_interval(1)[of "?x" a b] by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5062
  ultimately show ?th1 by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5063
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5064
  { 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
  5065
    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
  5066
    hence "a\<bullet>i \<le> b\<bullet>i" by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5067
    hence False using as by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5068
  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
  5069
  { assume as:"\<forall>i\<in>Basis. \<not> (b\<bullet>i < a\<bullet>i)"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5070
    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
  5071
    { 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
  5072
      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
  5073
      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
  5074
        by (auto simp: inner_add_left) }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5075
    hence "{a .. b} \<noteq> {}" using mem_interval(2)[of "?x" a b] by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5076
  ultimately show ?th2 by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5077
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5078
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5079
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
  5080
  "{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
  5081
  "{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
  5082
  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
  5083
44584
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  5084
lemma interval_sing:
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  5085
  fixes a :: "'a::ordered_euclidean_space"
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  5086
  shows "{a .. a} = {a}" and "{a<..<a} = {}"
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  5087
  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
  5088
  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
  5089
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5090
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
  5091
 "(\<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
  5092
 "(\<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
  5093
 "(\<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
  5094
 "(\<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
  5095
  unfolding subset_eq[unfolded Ball_def] unfolding mem_interval
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  5096
  by (best intro: order_trans less_le_trans le_less_trans less_imp_le)+
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  5097
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  5098
lemma interval_open_subset_closed:
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  5099
  fixes a :: "'a::ordered_euclidean_space"
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  5100
  shows "{a<..<b} \<subseteq> {a .. b}"
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  5101
  unfolding subset_eq [unfolded Ball_def] mem_interval
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  5102
  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
  5103
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5104
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
  5105
 "{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
  5106
 "{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
  5107
 "{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
  5108
 "{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
  5109
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5110
  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
  5111
  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
  5112
  { 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
  5113
    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
  5114
    fix i :: 'a assume i:"i\<in>Basis"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5115
    (** 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
  5116
    { 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
  5117
      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
  5118
      { 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
  5119
        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
  5120
          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
  5121
          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
  5122
      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
  5123
      moreover
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5124
      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
  5125
        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
  5126
        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
  5127
        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
  5128
      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
  5129
    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
  5130
    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
  5131
    { 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
  5132
      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
  5133
      { 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
  5134
        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
  5135
          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
  5136
          by (auto simp add: as2) }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5137
      hence "?x\<in>{c<..<d}" unfolding mem_interval by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5138
      moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5139
      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
  5140
        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
  5141
        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
  5142
        by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5143
      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
  5144
    hence "b\<bullet>i \<ge> d\<bullet>i" by(rule ccontr)auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5145
    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
  5146
    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
  5147
  } 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
  5148
  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
  5149
    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
  5150
    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
  5151
    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
  5152
    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
  5153
    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
  5154
    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
  5155
    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
  5156
  { 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
  5157
    fix i :: 'a assume i:"i\<in>Basis"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5158
    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
  5159
    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
  5160
  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
  5161
    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
  5162
    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
  5163
qed
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5164
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5165
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
  5166
 "{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
  5167
  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
  5168
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5169
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
  5170
  "{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
  5171
  "{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
  5172
  "{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
  5173
  "{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
  5174
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
  5175
  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
  5176
  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
  5177
      (\<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
  5178
    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
  5179
  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
  5180
  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
  5181
  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
  5182
  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
  5183
  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
  5184
qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5185
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5186
(* Moved interval_open_subset_closed a bit upwards *)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5187
44250
9133bc634d9c simplify proofs of lemmas open_interval, closed_interval
huffman
parents: 44233
diff changeset
  5188
lemma open_interval[intro]:
9133bc634d9c simplify proofs of lemmas open_interval, closed_interval
huffman
parents: 44233
diff changeset
  5189
  fixes a b :: "'a::ordered_euclidean_space" shows "open {a<..<b}"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5190
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
  5191
  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
  5192
    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
  5193
      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
  5194
  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
  5195
    by (auto simp add: eucl_less [where 'a='a])
9133bc634d9c simplify proofs of lemmas open_interval, closed_interval
huffman
parents: 44233
diff changeset
  5196
  finally show "open {a<..<b}" .
9133bc634d9c simplify proofs of lemmas open_interval, closed_interval
huffman
parents: 44233
diff changeset
  5197
qed
9133bc634d9c simplify proofs of lemmas open_interval, closed_interval
huffman
parents: 44233
diff changeset
  5198
9133bc634d9c simplify proofs of lemmas open_interval, closed_interval
huffman
parents: 44233
diff changeset
  5199
lemma closed_interval[intro]:
9133bc634d9c simplify proofs of lemmas open_interval, closed_interval
huffman
parents: 44233
diff changeset
  5200
  fixes a b :: "'a::ordered_euclidean_space" shows "closed {a .. b}"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5201
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
  5202
  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
  5203
    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
  5204
      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
  5205
  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
  5206
    by (auto simp add: eucl_le [where 'a='a])
9133bc634d9c simplify proofs of lemmas open_interval, closed_interval
huffman
parents: 44233
diff changeset
  5207
  finally show "closed {a .. b}" .
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5208
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5209
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  5210
lemma interior_closed_interval [intro]:
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  5211
  fixes a b :: "'a::ordered_euclidean_space"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  5212
  shows "interior {a..b} = {a<..<b}" (is "?L = ?R")
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5213
proof(rule subset_antisym)
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  5214
  show "?R \<subseteq> ?L" using interval_open_subset_closed open_interval
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  5215
    by (rule interior_maximal)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5216
next
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  5217
  { fix x assume "x \<in> interior {a..b}"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  5218
    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
  5219
    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
  5220
    { 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
  5221
      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
  5222
           "dist (x + (e / 2) *\<^sub>R i) x < e"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5223
        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
  5224
        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
  5225
      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
  5226
                     "(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
  5227
        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
  5228
        and   e[THEN spec[where x="x + (e/2) *\<^sub>R i"]]
44584
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  5229
        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
  5230
      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
  5231
        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
  5232
    hence "x \<in> {a<..<b}" unfolding mem_interval by auto  }
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44518
diff changeset
  5233
  thus "?L \<subseteq> ?R" ..
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5234
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5235
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5236
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
  5237
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
  5238
  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
  5239
  { 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
  5240
    { 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
  5241
      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
  5242
    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
  5243
    hence "norm x \<le> ?b" using norm_le_l1[of x] by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5244
  thus ?thesis unfolding interval and bounded_iff by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5245
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5246
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5247
lemma bounded_interval: fixes a :: "'a::ordered_euclidean_space" shows
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5248
 "bounded {a .. b} \<and> bounded {a<..<b}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5249
  using bounded_closed_interval[of a b]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5250
  using interval_open_subset_closed[of a b]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5251
  using bounded_subset[of "{a..b}" "{a<..<b}"]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5252
  by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5253
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5254
lemma not_interval_univ: fixes a :: "'a::ordered_euclidean_space" shows
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5255
 "({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
  5256
  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
  5257
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5258
lemma compact_interval: fixes a :: "'a::ordered_euclidean_space" shows "compact {a .. b}"
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5259
  using bounded_closed_imp_compact[of "{a..b}"] using bounded_interval[of a b]
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5260
  by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5261
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5262
lemma open_interval_midpoint: fixes a :: "'a::ordered_euclidean_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5263
  assumes "{a<..<b} \<noteq> {}" shows "((1/2) *\<^sub>R (a + b)) \<in> {a<..<b}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5264
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
  5265
  { 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
  5266
    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
  5267
      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
  5268
  thus ?thesis unfolding mem_interval by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5269
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5270
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5271
lemma open_closed_interval_convex: fixes x :: "'a::ordered_euclidean_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5272
  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
  5273
  shows "(e *\<^sub>R x + (1 - e) *\<^sub>R y) \<in> {a<..<b}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5274
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
  5275
  { 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
  5276
    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
  5277
    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
  5278
      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
  5279
      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
  5280
      using y unfolding mem_interval using i apply simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5281
      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
  5282
    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
  5283
    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
  5284
    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
  5285
    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
  5286
      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
  5287
      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
  5288
      using y unfolding mem_interval using i apply simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5289
      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
  5290
    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
  5291
    } 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
  5292
  thus ?thesis unfolding mem_interval by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5293
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5294
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5295
lemma closure_open_interval: fixes a :: "'a::ordered_euclidean_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5296
  assumes "{a<..<b} \<noteq> {}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5297
  shows "closure {a<..<b} = {a .. b}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5298
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
  5299
  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
  5300
  let ?c = "(1 / 2) *\<^sub>R (a + b)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5301
  { fix x assume as:"x \<in> {a .. b}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5302
    def f == "\<lambda>n::nat. x + (inverse (real n + 1)) *\<^sub>R (?c - x)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5303
    { 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
  5304
      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
  5305
      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
  5306
        x + (inverse (real n + 1)) *\<^sub>R (((1 / 2) *\<^sub>R (a + b)) - x)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5307
        by (auto simp add: algebra_simps)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5308
      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
  5309
      hence False using fn unfolding f_def using xc by auto  }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5310
    moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5311
    { assume "\<not> (f ---> x) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5312
      { fix e::real assume "e>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5313
        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
  5314
        then obtain N::nat where "inverse (real (N + 1)) < e" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5315
        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
  5316
        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
  5317
      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
  5318
        unfolding LIMSEQ_def by(auto simp add: dist_norm)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5319
      hence "(f ---> x) sequentially" unfolding f_def
44125
230a8665c919 mark some redundant theorems as legacy
huffman
parents: 44122
diff changeset
  5320
        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
  5321
        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
  5322
    ultimately have "x \<in> closure {a<..<b}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5323
      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
  5324
  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
  5325
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5326
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5327
lemma bounded_subset_open_interval_symmetric: fixes s::"('a::ordered_euclidean_space) set"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5328
  assumes "bounded s"  shows "\<exists>a. s \<subseteq> {-a<..<a}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5329
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5330
  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
  5331
  def a \<equiv> "(\<Sum>i\<in>Basis. (b + 1) *\<^sub>R i)::'a"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5332
  { 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
  5333
    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
  5334
    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
  5335
      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
  5336
  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
  5337
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5338
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5339
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
  5340
  fixes s :: "('a::ordered_euclidean_space) set"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5341
  shows "bounded s ==> (\<exists>a b. s \<subseteq> {a<..<b})"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5342
  by (auto dest!: bounded_subset_open_interval_symmetric)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5343
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5344
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
  5345
  fixes s :: "('a::ordered_euclidean_space) set"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5346
  assumes "bounded s" shows "\<exists>a. s \<subseteq> {-a .. a}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5347
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5348
  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
  5349
  thus ?thesis using interval_open_subset_closed[of "-a" a] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5350
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5351
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5352
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
  5353
  fixes s :: "('a::ordered_euclidean_space) set"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5354
  shows "bounded s ==> (\<exists>a b. s \<subseteq> {a .. b})"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5355
  using bounded_subset_closed_interval_symmetric[of s] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5356
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5357
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
  5358
  fixes a b :: "'a::ordered_euclidean_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5359
  shows "frontier {a .. b} = {a .. b} - {a<..<b}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5360
  unfolding frontier_def unfolding interior_closed_interval and closure_closed[OF closed_interval] ..
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5361
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5362
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
  5363
  fixes a b :: "'a::ordered_euclidean_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5364
  shows "frontier {a<..<b} = (if {a<..<b} = {} then {} else {a .. b} - {a<..<b})"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5365
proof(cases "{a<..<b} = {}")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5366
  case True thus ?thesis using frontier_empty by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5367
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5368
  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
  5369
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5370
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5371
lemma inter_interval_mixed_eq_empty: fixes a :: "'a::ordered_euclidean_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5372
  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
  5373
  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
  5374
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5375
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5376
(* Some stuff for half-infinite intervals too; FIXME: notation?  *)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5377
37673
f69f4b079275 generalize more lemmas from ordered_euclidean_space to euclidean_space
huffman
parents: 37649
diff changeset
  5378
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
  5379
  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
  5380
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
  5381
  { 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
  5382
    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
  5383
    { 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
  5384
      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
  5385
        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
  5386
      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
  5387
        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
  5388
    hence "x\<bullet>i \<le> b\<bullet>i" by(rule ccontr)auto  }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5389
  thus ?thesis unfolding closed_limpt unfolding islimpt_approachable by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5390
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5391
37673
f69f4b079275 generalize more lemmas from ordered_euclidean_space to euclidean_space
huffman
parents: 37649
diff changeset
  5392
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
  5393
  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
  5394
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
  5395
  { 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
  5396
    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
  5397
    { 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
  5398
      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
  5399
        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
  5400
      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
  5401
    hence "a\<bullet>i \<le> x\<bullet>i" by(rule ccontr)auto  }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5402
  thus ?thesis unfolding closed_limpt unfolding islimpt_approachable by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5403
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5404
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5405
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
  5406
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
  5407
  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
  5408
    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
  5409
  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
  5410
    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
  5411
  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
  5412
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
  5413
50881
ae630bab13da renamed countable_basis_space to second_countable_topology
hoelzl
parents: 50526
diff changeset
  5414
instance euclidean_space \<subseteq> second_countable_topology
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  5415
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
  5416
  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
  5417
  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
  5418
  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
  5419
  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
  5420
  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
  5421
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5422
  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
  5423
    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
  5424
  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
  5425
  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
  5426
  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
  5427
  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
  5428
    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
  5429
    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
  5430
      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
  5431
      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
  5432
      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
  5433
      done
50087
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  5434
  qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  5435
  ultimately
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50105
diff changeset
  5436
  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
  5437
qed
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  5438
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  5439
instance ordered_euclidean_space \<subseteq> polish_space ..
635d73673b5e regularity of measures, therefore:
immler
parents: 49962
diff changeset
  5440
36439
a65320184de9 move intervals section heading
huffman
parents: 36438
diff changeset
  5441
text {* Intervals in general, including infinite and mixtures of open and closed. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5442
37732
6432bf0d7191 generalize type of is_interval to class euclidean_space
huffman
parents: 37680
diff changeset
  5443
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
  5444
  (\<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
  5445
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5446
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
  5447
  "is_interval {a<..<b}" (is ?th2) proof -
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5448
  show ?th1 ?th2  unfolding is_interval_def mem_interval Ball_def atLeastAtMost_iff
44584
08ad27488983 simplify some proofs
huffman
parents: 44571
diff changeset
  5449
    by(meson order_trans le_less_trans less_le_trans less_trans)+ qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5450
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5451
lemma is_interval_empty:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5452
 "is_interval {}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5453
  unfolding is_interval_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5454
  by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5455
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5456
lemma is_interval_univ:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5457
 "is_interval UNIV"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5458
  unfolding is_interval_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5459
  by simp
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
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  5462
subsection {* Closure of halfspaces and hyperplanes *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5463
44219
f738e3200e24 generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
huffman
parents: 44216
diff changeset
  5464
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
  5465
  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
  5466
proof -
f738e3200e24 generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
huffman
parents: 44216
diff changeset
  5467
  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
  5468
    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
  5469
  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
  5470
    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
  5471
  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
  5472
    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
  5473
qed
f738e3200e24 generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
huffman
parents: 44216
diff changeset
  5474
f738e3200e24 generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
huffman
parents: 44216
diff changeset
  5475
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
  5476
  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
  5477
  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
  5478
  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
  5479
44213
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5480
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
  5481
  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
  5482
  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
  5483
  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
  5484
  shows "open {x. f x < g x}"
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5485
proof -
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5486
  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
  5487
    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
  5488
    by (rule isCont_open_vimage)
44213
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5489
  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
  5490
    by auto
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5491
  finally show ?thesis .
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5492
qed
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5493
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5494
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
  5495
  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
  5496
  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
  5497
  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
  5498
  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
  5499
proof -
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5500
  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
  5501
    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
  5502
    by (rule isCont_closed_vimage)
44213
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5503
  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
  5504
    by auto
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5505
  finally show ?thesis .
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5506
qed
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5507
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5508
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
  5509
  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
  5510
  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
  5511
  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
  5512
  shows "closed {x. f x = g x}"
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5513
proof -
44216
903bfe95fece generalized lemma closed_Collect_eq
huffman
parents: 44213
diff changeset
  5514
  have "open {(x::'b, y::'b). x \<noteq> y}"
903bfe95fece generalized lemma closed_Collect_eq
huffman
parents: 44213
diff changeset
  5515
    unfolding open_prod_def by (auto dest!: hausdorff)
903bfe95fece generalized lemma closed_Collect_eq
huffman
parents: 44213
diff changeset
  5516
  hence "closed {(x::'b, y::'b). x = y}"
903bfe95fece generalized lemma closed_Collect_eq
huffman
parents: 44213
diff changeset
  5517
    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
  5518
  with isCont_Pair [OF f g]
44216
903bfe95fece generalized lemma closed_Collect_eq
huffman
parents: 44213
diff changeset
  5519
  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
  5520
    by (rule isCont_closed_vimage)
44216
903bfe95fece generalized lemma closed_Collect_eq
huffman
parents: 44213
diff changeset
  5521
  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
  5522
  finally show ?thesis .
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5523
qed
6fb54701a11b add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents: 44212
diff changeset
  5524
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5525
lemma continuous_at_inner: "continuous (at x) (inner a)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5526
  unfolding continuous_at by (intro tendsto_intros)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5527
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5528
lemma closed_halfspace_le: "closed {x. inner a x \<le> b}"
44233
aa74ce315bae add simp rules for isCont
huffman
parents: 44219
diff changeset
  5529
  by (simp add: closed_Collect_le)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5530
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5531
lemma closed_halfspace_ge: "closed {x. inner a x \<ge> b}"
44233
aa74ce315bae add simp rules for isCont
huffman
parents: 44219
diff changeset
  5532
  by (simp add: closed_Collect_le)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5533
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5534
lemma closed_hyperplane: "closed {x. inner a x = b}"
44233
aa74ce315bae add simp rules for isCont
huffman
parents: 44219
diff changeset
  5535
  by (simp add: closed_Collect_eq)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5536
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5537
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
  5538
  shows "closed {x::'a::euclidean_space. x\<bullet>i \<le> a}"
44233
aa74ce315bae add simp rules for isCont
huffman
parents: 44219
diff changeset
  5539
  by (simp add: closed_Collect_le)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5540
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5541
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
  5542
  shows "closed {x::'a::euclidean_space. x\<bullet>i \<ge> a}"
44233
aa74ce315bae add simp rules for isCont
huffman
parents: 44219
diff changeset
  5543
  by (simp add: closed_Collect_le)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5544
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  5545
text {* Openness of halfspaces. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5546
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5547
lemma open_halfspace_lt: "open {x. inner a x < b}"
44233
aa74ce315bae add simp rules for isCont
huffman
parents: 44219
diff changeset
  5548
  by (simp add: open_Collect_less)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5549
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5550
lemma open_halfspace_gt: "open {x. inner a x > b}"
44233
aa74ce315bae add simp rules for isCont
huffman
parents: 44219
diff changeset
  5551
  by (simp add: open_Collect_less)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5552
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5553
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
  5554
  shows "open {x::'a::euclidean_space. x\<bullet>i < a}"
44233
aa74ce315bae add simp rules for isCont
huffman
parents: 44219
diff changeset
  5555
  by (simp add: open_Collect_less)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5556
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5557
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
  5558
  shows "open {x::'a::euclidean_space. x\<bullet>i > a}"
44233
aa74ce315bae add simp rules for isCont
huffman
parents: 44219
diff changeset
  5559
  by (simp add: open_Collect_less)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5560
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5561
text{* Instantiation for intervals on @{text ordered_euclidean_space} *}
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5562
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5563
lemma eucl_lessThan_eq_halfspaces:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5564
  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
  5565
  shows "{..<a} = (\<Inter>i\<in>Basis. {x. x \<bullet> i < a \<bullet> i})"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5566
 by (auto simp: eucl_less[where 'a='a])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5567
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5568
lemma eucl_greaterThan_eq_halfspaces:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5569
  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
  5570
  shows "{a<..} = (\<Inter>i\<in>Basis. {x. a \<bullet> i < x \<bullet> i})"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5571
 by (auto simp: eucl_less[where 'a='a])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5572
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5573
lemma eucl_atMost_eq_halfspaces:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5574
  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
  5575
  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
  5576
 by (auto simp: eucl_le[where 'a='a])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5577
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5578
lemma eucl_atLeast_eq_halfspaces:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5579
  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
  5580
  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
  5581
 by (auto simp: eucl_le[where 'a='a])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5582
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5583
lemma open_eucl_lessThan[simp, intro]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5584
  fixes a :: "'a\<Colon>ordered_euclidean_space"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5585
  shows "open {..< a}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5586
  by (auto simp: eucl_lessThan_eq_halfspaces open_halfspace_component_lt)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5587
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5588
lemma open_eucl_greaterThan[simp, intro]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5589
  fixes a :: "'a\<Colon>ordered_euclidean_space"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5590
  shows "open {a <..}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5591
  by (auto simp: eucl_greaterThan_eq_halfspaces open_halfspace_component_gt)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5592
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5593
lemma closed_eucl_atMost[simp, intro]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5594
  fixes a :: "'a\<Colon>ordered_euclidean_space"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5595
  shows "closed {.. a}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5596
  unfolding eucl_atMost_eq_halfspaces
44233
aa74ce315bae add simp rules for isCont
huffman
parents: 44219
diff changeset
  5597
  by (simp add: closed_INT closed_Collect_le)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5598
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5599
lemma closed_eucl_atLeast[simp, intro]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5600
  fixes a :: "'a\<Colon>ordered_euclidean_space"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5601
  shows "closed {a ..}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5602
  unfolding eucl_atLeast_eq_halfspaces
44233
aa74ce315bae add simp rules for isCont
huffman
parents: 44219
diff changeset
  5603
  by (simp add: closed_INT closed_Collect_le)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
  5604
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  5605
text {* This gives a simple derivation of limit component bounds. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5606
37673
f69f4b079275 generalize more lemmas from ordered_euclidean_space to euclidean_space
huffman
parents: 37649
diff changeset
  5607
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
  5608
  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
  5609
  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
  5610
  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
  5611
37673
f69f4b079275 generalize more lemmas from ordered_euclidean_space to euclidean_space
huffman
parents: 37649
diff changeset
  5612
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
  5613
  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
  5614
  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
  5615
  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
  5616
37673
f69f4b079275 generalize more lemmas from ordered_euclidean_space to euclidean_space
huffman
parents: 37649
diff changeset
  5617
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
  5618
  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
  5619
  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
  5620
  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
  5621
  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
  5622
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5623
text{* Limits relative to a union.                                               *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5624
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5625
lemma eventually_within_Un:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5626
  "eventually P (net within (s \<union> t)) \<longleftrightarrow>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5627
    eventually P (net within s) \<and> eventually P (net within t)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5628
  unfolding Limits.eventually_within
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5629
  by (auto elim!: eventually_rev_mp)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5630
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5631
lemma Lim_within_union:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5632
 "(f ---> l) (net within (s \<union> t)) \<longleftrightarrow>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5633
  (f ---> l) (net within s) \<and> (f ---> l) (net within t)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5634
  unfolding tendsto_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5635
  by (auto simp add: eventually_within_Un)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5636
36442
b96d9dc6acca generalize more continuity lemmas
huffman
parents: 36441
diff changeset
  5637
lemma Lim_topological:
b96d9dc6acca generalize more continuity lemmas
huffman
parents: 36441
diff changeset
  5638
 "(f ---> l) net \<longleftrightarrow>
b96d9dc6acca generalize more continuity lemmas
huffman
parents: 36441
diff changeset
  5639
        trivial_limit net \<or>
b96d9dc6acca generalize more continuity lemmas
huffman
parents: 36441
diff changeset
  5640
        (\<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
  5641
  unfolding tendsto_def trivial_limit_eq by auto
b96d9dc6acca generalize more continuity lemmas
huffman
parents: 36441
diff changeset
  5642
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5643
lemma continuous_on_union:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5644
  assumes "closed s" "closed t" "continuous_on s f" "continuous_on t f"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5645
  shows "continuous_on (s \<union> t) f"
36442
b96d9dc6acca generalize more continuity lemmas
huffman
parents: 36441
diff changeset
  5646
  using assms unfolding continuous_on Lim_within_union
b96d9dc6acca generalize more continuity lemmas
huffman
parents: 36441
diff changeset
  5647
  unfolding Lim_topological trivial_limit_within closed_limpt by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5648
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5649
lemma continuous_on_cases:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5650
  assumes "closed s" "closed t" "continuous_on s f" "continuous_on t g"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5651
          "\<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
  5652
  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
  5653
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5654
  let ?h = "(\<lambda>x. if P x then f x else g x)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5655
  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
  5656
  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
  5657
  moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5658
  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
  5659
  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
  5660
  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
  5661
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5662
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5663
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5664
text{* Some more convenient intermediate-value theorem formulations.             *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5665
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5666
lemma connected_ivt_hyperplane:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5667
  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
  5668
  shows "\<exists>z \<in> s. inner a z = b"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5669
proof(rule ccontr)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5670
  assume as:"\<not> (\<exists>z\<in>s. inner a z = b)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5671
  let ?A = "{x. inner a x < b}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5672
  let ?B = "{x. inner a x > b}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5673
  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
  5674
  moreover have "?A \<inter> ?B = {}" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5675
  moreover have "s \<subseteq> ?A \<union> ?B" using as by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5676
  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
  5677
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5678
37673
f69f4b079275 generalize more lemmas from ordered_euclidean_space to euclidean_space
huffman
parents: 37649
diff changeset
  5679
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
  5680
 "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
  5681
  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
  5682
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  5683
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  5684
subsection {* Homeomorphisms *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5685
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5686
definition "homeomorphism s t f g \<equiv>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5687
     (\<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
  5688
     (\<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
  5689
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5690
definition
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5691
  homeomorphic :: "'a::metric_space set \<Rightarrow> 'b::metric_space set \<Rightarrow> bool"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5692
    (infixr "homeomorphic" 60) where
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5693
  homeomorphic_def: "s homeomorphic t \<equiv> (\<exists>f g. homeomorphism s t f g)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5694
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5695
lemma homeomorphic_refl: "s homeomorphic s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5696
  unfolding homeomorphic_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5697
  unfolding homeomorphism_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5698
  using continuous_on_id
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5699
  apply(rule_tac x = "(\<lambda>x. x)" in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5700
  apply(rule_tac x = "(\<lambda>x. x)" in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5701
  by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5702
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5703
lemma homeomorphic_sym:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5704
 "s homeomorphic t \<longleftrightarrow> t homeomorphic s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5705
unfolding homeomorphic_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5706
unfolding homeomorphism_def
33324
51eb2ffa2189 Tidied up some very ugly proofs
paulson
parents: 33270
diff changeset
  5707
by blast 
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5708
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5709
lemma homeomorphic_trans:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5710
  assumes "s homeomorphic t" "t homeomorphic u" shows "s homeomorphic u"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5711
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5712
  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
  5713
    using assms(1) unfolding homeomorphic_def homeomorphism_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5714
  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
  5715
    using assms(2) unfolding homeomorphic_def homeomorphism_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5716
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5717
  { 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
  5718
  moreover have "(f2 \<circ> f1) ` s = u" using fg1(2) fg2(2) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5719
  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
  5720
  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
  5721
  moreover have "(g1 \<circ> g2) ` u = s" using fg1(5) fg2(5) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5722
  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
  5723
  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
  5724
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5725
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5726
lemma homeomorphic_minimal:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5727
 "s homeomorphic t \<longleftrightarrow>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5728
    (\<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
  5729
           (\<forall>y\<in>t. g(y) \<in> s \<and> (f(g(y)) = y)) \<and>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5730
           continuous_on s f \<and> continuous_on t g)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5731
unfolding homeomorphic_def homeomorphism_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5732
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
  5733
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
  5734
unfolding image_iff
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5735
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
  5736
apply auto apply(rule_tac x="g x" in bexI) apply auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5737
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
  5738
apply auto apply(rule_tac x="f x" in bexI) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5739
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  5740
text {* Relatively weak hypotheses if a set is compact. *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5741
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5742
lemma homeomorphism_compact:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5743
  fixes f :: "'a::heine_borel \<Rightarrow> 'b::heine_borel"
44647
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  5744
    (* class constraint due to continuous_on_inv *)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5745
  assumes "compact s" "continuous_on s f"  "f ` s = t"  "inj_on f s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5746
  shows "\<exists>g. homeomorphism s t f g"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5747
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5748
  def g \<equiv> "\<lambda>x. SOME y. y\<in>s \<and> f y = x"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5749
  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
  5750
  { fix y assume "y\<in>t"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5751
    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
  5752
    hence "g (f x) = x" using g by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5753
    hence "f (g y) = y" unfolding x(1)[THEN sym] by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5754
  hence g':"\<forall>x\<in>t. f (g x) = x" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5755
  moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5756
  { fix x
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5757
    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
  5758
    moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5759
    { assume "x\<in>g ` t"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5760
      then obtain y where y:"y\<in>t" "g y = x" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5761
      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
  5762
      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
  5763
    ultimately have "x\<in>s \<longleftrightarrow> x \<in> g ` t" ..  }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5764
  hence "g ` t = s" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5765
  ultimately
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5766
  show ?thesis unfolding homeomorphism_def homeomorphic_def
44647
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  5767
    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
  5768
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5769
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5770
lemma homeomorphic_compact:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5771
  fixes f :: "'a::heine_borel \<Rightarrow> 'b::heine_borel"
44647
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44632
diff changeset
  5772
    (* class constraint due to continuous_on_inv *)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5773
  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
  5774
          \<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
  5775
  unfolding homeomorphic_def by (metis homeomorphism_compact)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5776
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5777
text{* Preservation of topological properties.                                   *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5778
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5779
lemma homeomorphic_compactness:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5780
 "s homeomorphic t ==> (compact s \<longleftrightarrow> compact t)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5781
unfolding homeomorphic_def homeomorphism_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5782
by (metis compact_continuous_image)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5783
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5784
text{* Results on translation, scaling etc.                                      *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5785
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5786
lemma homeomorphic_scaling:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5787
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5788
  assumes "c \<noteq> 0"  shows "s homeomorphic ((\<lambda>x. c *\<^sub>R x) ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5789
  unfolding homeomorphic_minimal
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5790
  apply(rule_tac x="\<lambda>x. c *\<^sub>R x" in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5791
  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
  5792
  using assms by (auto simp add: continuous_on_intros)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5793
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5794
lemma homeomorphic_translation:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5795
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5796
  shows "s homeomorphic ((\<lambda>x. a + x) ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5797
  unfolding homeomorphic_minimal
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5798
  apply(rule_tac x="\<lambda>x. a + x" in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5799
  apply(rule_tac x="\<lambda>x. -a + x" in exI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5800
  using continuous_on_add[OF continuous_on_const continuous_on_id] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5801
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5802
lemma homeomorphic_affinity:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5803
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5804
  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
  5805
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5806
  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
  5807
  show ?thesis
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5808
    using homeomorphic_trans
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5809
    using homeomorphic_scaling[OF assms, of s]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5810
    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
  5811
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5812
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5813
lemma homeomorphic_balls:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5814
  fixes a b ::"'a::real_normed_vector" (* FIXME: generalize to metric_space *)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5815
  assumes "0 < d"  "0 < e"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5816
  shows "(ball a d) homeomorphic  (ball b e)" (is ?th)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5817
        "(cball a d) homeomorphic (cball b e)" (is ?cth)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5818
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5819
  have *:"\<bar>e / d\<bar> > 0" "\<bar>d / e\<bar> >0" using assms using divide_pos_pos by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5820
  show ?th unfolding homeomorphic_minimal
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5821
    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
  5822
    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
  5823
    using assms apply (auto simp add: dist_commute)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5824
    unfolding dist_norm
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5825
    apply (auto simp add: pos_divide_less_eq mult_strict_left_mono)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5826
    unfolding continuous_on
36659
f794e92784aa adapt to removed premise on tendsto lemma (cf. 88f0125c3bd2)
huffman
parents: 36623
diff changeset
  5827
    by (intro ballI tendsto_intros, simp)+
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5828
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5829
  have *:"\<bar>e / d\<bar> > 0" "\<bar>d / e\<bar> >0" using assms using divide_pos_pos by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5830
  show ?cth unfolding homeomorphic_minimal
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5831
    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
  5832
    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
  5833
    using assms apply (auto simp add: dist_commute)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5834
    unfolding dist_norm
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5835
    apply (auto simp add: pos_divide_le_eq)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5836
    unfolding continuous_on
36659
f794e92784aa adapt to removed premise on tendsto lemma (cf. 88f0125c3bd2)
huffman
parents: 36623
diff changeset
  5837
    by (intro ballI tendsto_intros, simp)+
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5838
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5839
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5840
text{* "Isometry" (up to constant bounds) of injective linear map etc.           *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5841
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5842
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
  5843
  fixes x :: "nat \<Rightarrow> 'a::euclidean_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5844
  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
  5845
  shows "Cauchy x"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5846
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5847
  interpret f: bounded_linear f by fact
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5848
  { fix d::real assume "d>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5849
    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
  5850
      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
  5851
    { 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
  5852
      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
  5853
        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
  5854
        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
  5855
      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
  5856
        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
  5857
      finally have "norm (x n - x N) < d" using `e>0` by simp }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5858
    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
  5859
  thus ?thesis unfolding cauchy and dist_norm by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5860
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5861
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5862
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
  5863
  fixes f :: "'a::euclidean_space => 'b::euclidean_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5864
  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
  5865
  shows "complete(f ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5866
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5867
  { 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
  5868
    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
  5869
      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
  5870
    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
  5871
    hence "f \<circ> x = g" unfolding fun_eq_iff by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5872
    then obtain l where "l\<in>s" and l:"(x ---> l) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5873
      using cs[unfolded complete_def, THEN spec[where x="x"]]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5874
      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
  5875
    hence "\<exists>l\<in>f ` s. (g ---> l) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5876
      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
  5877
      unfolding `f \<circ> x = g` by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5878
  thus ?thesis unfolding complete_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5879
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5880
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5881
lemma dist_0_norm:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5882
  fixes x :: "'a::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5883
  shows "dist 0 x = norm x"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5884
unfolding dist_norm by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5885
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  5886
lemma injective_imp_isometric: fixes f::"'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5887
  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
  5888
  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
  5889
proof(cases "s \<subseteq> {0::'a}")
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5890
  case True
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5891
  { fix x assume "x \<in> s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5892
    hence "x = 0" using True by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5893
    hence "norm x \<le> norm (f x)" by auto  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5894
  thus ?thesis by(auto intro!: exI[where x=1])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5895
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5896
  interpret f: bounded_linear f by fact
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5897
  case False
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5898
  then obtain a where a:"a\<noteq>0" "a\<in>s" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5899
  from False have "s \<noteq> {}" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5900
  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
  5901
  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
  5902
  let ?S'' = "{x::'a. norm x = norm a}"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5903
36362
06475a1547cb fix lots of looping simp calls and other warnings
huffman
parents: 36360
diff changeset
  5904
  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
  5905
  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
  5906
  moreover have "?S' = s \<inter> ?S''" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5907
  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
  5908
  moreover have *:"f ` ?S' = ?S" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5909
  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
  5910
  hence "closed ?S" using compact_imp_closed by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5911
  moreover have "?S \<noteq> {}" using a by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5912
  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
  5913
  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
  5914
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5915
  let ?e = "norm (f b) / norm b"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5916
  have "norm b > 0" using ba and a and norm_ge_zero by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5917
  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
  5918
  ultimately have "0 < norm (f b) / norm b" by(simp only: divide_pos_pos)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5919
  moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5920
  { fix x assume "x\<in>s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5921
    hence "norm (f b) / norm b * norm x \<le> norm (f x)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5922
    proof(cases "x=0")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5923
      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
  5924
    next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5925
      case False
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5926
      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
  5927
      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
  5928
      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
  5929
      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
  5930
        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
  5931
        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
  5932
    qed }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5933
  ultimately
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5934
  show ?thesis by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5935
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5936
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5937
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
  5938
  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5939
  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
  5940
  shows "closed(f ` s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5941
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5942
  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
  5943
  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
  5944
    unfolding complete_eq_closed[THEN sym] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5945
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5946
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  5947
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  5948
subsection {* Some properties of a canonical subspace *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5949
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5950
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
  5951
  "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
  5952
  unfolding subspace_def by (auto simp: inner_add_left)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5953
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5954
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
  5955
 "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
  5956
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
  5957
  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
  5958
  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
  5959
    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
  5960
  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
  5961
    by auto
d366fa5551ef declare euclidean_simps [simp] at the point they are proved;
huffman
parents: 44365
diff changeset
  5962
  finally show "closed ?A" .
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5963
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5964
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5965
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
  5966
  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
  5967
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
  5968
  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
  5969
  have "d \<subseteq> ?A" using d by (auto simp: inner_Basis)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5970
  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
  5971
  { 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
  5972
    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
  5973
    from this d have "x \<in> span d"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5974
    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
  5975
      case empty hence "x=0" apply(rule_tac euclidean_eqI) by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5976
      thus ?case using subspace_0[OF subspace_span[of "{}"]] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5977
    next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5978
      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
  5979
      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
  5980
      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
  5981
      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
  5982
      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
  5983
      { 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
  5984
        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
  5985
          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
  5986
      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
  5987
      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
  5988
        using span_mono[of F "insert k F"] using assms by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5989
      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
  5990
      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
  5991
      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
  5992
        using span_mul by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5993
      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
  5994
      have "y + (x\<bullet>k) *\<^sub>R k \<in> span (insert k F)"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5995
        using span_add by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5996
      thus ?case using y by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5997
    qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5998
  }
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  5999
  hence "?A \<subseteq> span d" by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6000
  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
  6001
  { 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
  6002
  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
  6003
  moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6004
  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
  6005
  ultimately show ?thesis using dim_unique[of d ?A] by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6006
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6007
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6008
text{* Hence closure and completeness of all subspaces.                          *}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6009
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50324
diff changeset
  6010
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
  6011
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
  6012
  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
  6013
  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
  6014
  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
  6015
    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
  6016
  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
  6017
    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
  6018
  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
  6019
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
  6020
  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
  6021
qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6022
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  6023
lemma closed_subspace: fixes s::"('a::euclidean_space) set"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6024
  assumes "subspace s" shows "closed s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6025
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
  6026
  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
  6027
  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
  6028
  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
  6029
  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
  6030
      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
  6031
    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
  6032
    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
  6033
  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
  6034
  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
  6035
  { 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
  6036
  moreover have "closed ?t" using closed_substandard .
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6037
  moreover have "subspace ?t" using subspace_substandard .
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6038
  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
  6039
    unfolding f(2) using f(1) unfolding linear_conv_bounded_linear by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6040
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6041
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6042
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
  6043
  fixes s :: "('a::euclidean_space) set" shows "subspace s ==> complete s"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6044
  using complete_eq_closed closed_subspace
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6045
  by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6046
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6047
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
  6048
  fixes s :: "('a::euclidean_space) set"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6049
  shows "dim(closure s) = dim s" (is "?dc = ?d")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6050
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6051
  have "?dc \<le> ?d" using closure_minimal[OF span_inc, of s]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6052
    using closed_subspace[OF subspace_span, of s]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6053
    using dim_subset[of "closure s" "span s"] unfolding dim_span by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6054
  thus ?thesis using dim_subset[OF closure_subset, of s] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6055
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6056
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  6057
36437
e76cb1d4663c reorganize subsection headings
huffman
parents: 36431
diff changeset
  6058
subsection {* Affine transformations of intervals *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6059
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6060
lemma real_affinity_le:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34999
diff changeset
  6061
 "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
  6062
  by (simp add: field_simps inverse_eq_divide)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6063
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6064
lemma real_le_affinity:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34999
diff changeset
  6065
 "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
  6066
  by (simp add: field_simps inverse_eq_divide)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6067
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6068
lemma real_affinity_lt:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34999
diff changeset
  6069
 "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
  6070
  by (simp add: field_simps inverse_eq_divide)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6071
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6072
lemma real_lt_affinity:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34999
diff changeset
  6073
 "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
  6074
  by (simp add: field_simps inverse_eq_divide)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6075
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6076
lemma real_affinity_eq:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34999
diff changeset
  6077
 "(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
  6078
  by (simp add: field_simps inverse_eq_divide)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6079
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6080
lemma real_eq_affinity:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34999
diff changeset
  6081
 "(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
  6082
  by (simp add: field_simps inverse_eq_divide)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6083
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6084
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
  6085
  fixes a b c :: "'a::ordered_euclidean_space"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6086
  shows "(\<lambda>x. m *\<^sub>R x + c) ` {a .. b} =
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6087
            (if {a .. b} = {} then {}
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6088
            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
  6089
            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
  6090
proof(cases "m=0")  
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6091
  { 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
  6092
    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
  6093
      apply(subst euclidean_eq_iff) by (auto intro: order_antisym) }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6094
  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
  6095
  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
  6096
  ultimately show ?thesis by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6097
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6098
  case False
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6099
  { fix y assume "a \<le> y" "y \<le> b" "m > 0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6100
    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
  6101
      unfolding eucl_le[where 'a='a] by (auto simp: inner_simps)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6102
  } moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6103
  { fix y assume "a \<le> y" "y \<le> b" "m < 0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6104
    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
  6105
      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
  6106
  } moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6107
  { 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
  6108
    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
  6109
      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
  6110
      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
  6111
      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
  6112
  } moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6113
  { 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
  6114
    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
  6115
      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
  6116
      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
  6117
      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
  6118
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6119
  ultimately show ?thesis using False by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6120
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6121
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 37452
diff changeset
  6122
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
  6123
  (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
  6124
  using image_affinity_interval[of m 0 a b] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6125
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  6126
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  6127
subsection {* Banach fixed point theorem (not really topological...) *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6128
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6129
lemma banach_fix:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6130
  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
  6131
          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
  6132
  shows "\<exists>! x\<in>s. (f x = x)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6133
proof-
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6134
  have "1 - c > 0" using c by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6135
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6136
  from s(2) obtain z0 where "z0 \<in> s" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6137
  def z \<equiv> "\<lambda>n. (f ^^ n) z0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6138
  { fix n::nat
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6139
    have "z n \<in> s" unfolding z_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6140
    proof(induct n) case 0 thus ?case using `z0 \<in>s` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6141
    next case Suc thus ?case using f by auto qed }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6142
  note z_in_s = this
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6143
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6144
  def d \<equiv> "dist (z 0) (z 1)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6145
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6146
  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
  6147
  { fix n::nat
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6148
    have "dist (z n) (z (Suc n)) \<le> (c ^ n) * d"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6149
    proof(induct n)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6150
      case 0 thus ?case unfolding d_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6151
    next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6152
      case (Suc m)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6153
      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
  6154
        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
  6155
      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
  6156
        unfolding fzn and mult_le_cancel_left by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6157
    qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6158
  } note cf_z = this
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6159
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6160
  { fix n m::nat
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6161
    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
  6162
    proof(induct n)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6163
      case 0 show ?case by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6164
    next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6165
      case (Suc k)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6166
      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
  6167
        using dist_triangle and c by(auto simp add: dist_triangle)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6168
      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
  6169
        using cf_z[of "m + k"] and c by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6170
      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
  6171
        using Suc by (auto simp add: field_simps)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6172
      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
  6173
        unfolding power_add by (auto simp add: field_simps)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6174
      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
  6175
        using c by (auto simp add: field_simps)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6176
      finally show ?case by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6177
    qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6178
  } note cf_z2 = this
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6179
  { fix e::real assume "e>0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6180
    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
  6181
    proof(cases "d = 0")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6182
      case True
41863
e5104b436ea1 removed dependency on Dense_Linear_Order
boehmes
parents: 41413
diff changeset
  6183
      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
  6184
        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
  6185
      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
  6186
        by (simp add: *)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6187
      thus ?thesis using `e>0` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6188
    next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6189
      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
  6190
        by (metis False d_def less_le)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6191
      hence "0 < e * (1 - c) / d" using `e>0` and `1-c>0`
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6192
        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
  6193
      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
  6194
      { 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
  6195
        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
  6196
        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
  6197
        hence **:"d * (1 - c ^ (m - n)) / (1 - c) > 0"
36778
739a9379e29b avoid using real-specific versions of generic lemmas
huffman
parents: 36669
diff changeset
  6198
          using mult_pos_pos[OF `d>0`, of "1 - c ^ (m - n)"]
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6199
          using divide_pos_pos[of "d * (1 - c ^ (m - n))" "1 - c"]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6200
          using `0 < 1 - c` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6201
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6202
        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
  6203
          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
  6204
          by (auto simp add: mult_commute dist_commute)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6205
        also have "\<dots> \<le> c ^ N * d * (1 - c ^ (m - n)) / (1 - c)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6206
          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
  6207
          unfolding mult_assoc by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6208
        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
  6209
          using mult_strict_right_mono[OF N **] unfolding mult_assoc by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6210
        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
  6211
        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
  6212
        finally have  "dist (z m) (z n) < e" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6213
      } note * = this
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6214
      { fix m n::nat assume as:"N\<le>m" "N\<le>n"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6215
        hence "dist (z n) (z m) < e"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6216
        proof(cases "n = m")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6217
          case True thus ?thesis using `e>0` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6218
        next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6219
          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
  6220
        qed }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6221
      thus ?thesis by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6222
    qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6223
  }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6224
  hence "Cauchy z" unfolding cauchy_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6225
  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
  6226
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6227
  def e \<equiv> "dist (f x) x"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6228
  have "e = 0" proof(rule ccontr)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6229
    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
  6230
      by (metis dist_eq_0_iff dist_nz e_def)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6231
    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
  6232
      using x[unfolded LIMSEQ_def, THEN spec[where x="e/2"]] by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6233
    hence N':"dist (z N) x < e / 2" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6234
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6235
    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
  6236
      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
  6237
      by (metis dist_eq_0_iff dist_nz order_less_asym less_le)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6238
    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
  6239
      using z_in_s[of N] `x\<in>s` using c by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6240
    also have "\<dots> < e / 2" using N' and c using * by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6241
    finally show False unfolding fzn
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6242
      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
  6243
      unfolding e_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6244
  qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6245
  hence "f x = x" unfolding e_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6246
  moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6247
  { fix y assume "f y = y" "y\<in>s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6248
    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
  6249
      using `x\<in>s` and `f x = x` by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6250
    hence "dist x y = 0" unfolding mult_le_cancel_right1
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6251
      using c and zero_le_dist[of x y] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6252
    hence "y = x" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6253
  }
34999
5312d2ffee3b Changed 'bounded unique existential quantifiers' from a constant to syntax translation.
hoelzl
parents: 34964
diff changeset
  6254
  ultimately show ?thesis using `x\<in>s` by blast+
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6255
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6256
44210
eba74571833b Topology_Euclidean_Space.thy: organize section headings
huffman
parents: 44207
diff changeset
  6257
subsection {* Edelstein fixed point theorem *}
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6258
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6259
lemma edelstein_fix:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6260
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6261
  assumes s:"compact s" "s \<noteq> {}" and gs:"(g ` s) \<subseteq> s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6262
      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
  6263
  shows "\<exists>! x\<in>s. g x = x"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6264
proof(cases "\<exists>x\<in>s. g x \<noteq> x")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6265
  obtain x where "x\<in>s" using s(2) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6266
  case False hence g:"\<forall>x\<in>s. g x = x" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6267
  { fix y assume "y\<in>s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6268
    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
  6269
      unfolding g[THEN bspec[where x=x], OF `x\<in>s`]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6270
      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
  6271
  thus ?thesis using `x\<in>s` and g by blast+
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6272
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6273
  case True
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6274
  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
  6275
  { fix x y assume "x \<in> s" "y \<in> s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6276
    hence "dist (g x) (g y) \<le> dist x y"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6277
      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
  6278
  def y \<equiv> "g x"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6279
  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
  6280
  def f \<equiv> "\<lambda>n. g ^^ n"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6281
  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
  6282
  have [simp]:"\<And>z. f 0 z = z" unfolding f_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6283
  { fix n::nat and z assume "z\<in>s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6284
    have "f n z \<in> s" unfolding f_def
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6285
    proof(induct n)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6286
      case 0 thus ?case using `z\<in>s` by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6287
    next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6288
      case (Suc n) thus ?case using gs[unfolded image_subset_iff] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6289
    qed } note fs = this
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6290
  { fix m n ::nat assume "m\<le>n"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6291
    fix w z assume "w\<in>s" "z\<in>s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6292
    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
  6293
    proof(induct n)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6294
      case 0 thus ?case by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6295
    next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6296
      case (Suc n)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6297
      thus ?case proof(cases "m\<le>n")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6298
        case True thus ?thesis using Suc(1)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6299
          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
  6300
      next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6301
        case False hence mn:"m = Suc n" using Suc(2) by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6302
        show ?thesis unfolding mn  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
    qed } note distf = this
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6305
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6306
  def h \<equiv> "\<lambda>n. (f n x, f n y)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6307
  let ?s2 = "s \<times> s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6308
  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
  6309
    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
  6310
    using fs[OF `x\<in>s`] and fs[OF `y\<in>s`] by blast
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6311
  def a \<equiv> "fst l" def b \<equiv> "snd l"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6312
  have lab:"l = (a, b)" unfolding a_def b_def by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6313
  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
  6314
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6315
  have lima:"((fst \<circ> (h \<circ> r)) ---> a) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6316
   and limb:"((snd \<circ> (h \<circ> r)) ---> b) sequentially"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6317
    using lr
44167
e81d676d598e avoid duplicate rule warnings
huffman
parents: 44139
diff changeset
  6318
    unfolding o_def a_def b_def by (rule tendsto_intros)+
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6319
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6320
  { fix n::nat
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6321
    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
  6322
    { fix x y :: 'a
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6323
      have "dist (-x) (-y) = dist x y" unfolding dist_norm
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6324
        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
  6325
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6326
    { assume as:"dist a b > dist (f n x) (f n y)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6327
      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
  6328
        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
  6329
        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
  6330
      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
  6331
        apply(erule_tac x="Na+Nb+n" in allE)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6332
        apply(erule_tac x="Na+Nb+n" in allE) apply simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6333
        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
  6334
          "-b"  "- f (r (Na + Nb + n)) y"]
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36336
diff changeset
  6335
        unfolding ** by (auto simp add: algebra_simps dist_commute)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6336
      moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6337
      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
  6338
        using distf[of n "r (Na+Nb+n)", OF _ `x\<in>s` `y\<in>s`]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6339
        using subseq_bigger[OF r, of "Na+Nb+n"]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6340
        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
  6341
      ultimately have False by simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6342
    }
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6343
    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
  6344
  note ab_fn = this
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6345
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6346
  have [simp]:"a = b" proof(rule ccontr)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6347
    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
  6348
    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
  6349
    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
  6350
      using lima limb unfolding LIMSEQ_def
93943da0a010 remove redundant lemma Lim_sequentially in favor of lemma LIMSEQ_def
huffman
parents: 44905
diff changeset
  6351
      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
  6352
    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
  6353
    have "dist (f (Suc n) x) (g a) \<le> dist (f n x) a"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6354
      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
  6355
    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
  6356
      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
  6357
    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
  6358
    thus False unfolding e_def using ab_fn[of "Suc n"] by norm
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
  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
  6362
  { fix x y assume "x\<in>s" "y\<in>s" moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6363
    fix e::real assume "e>0" ultimately
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 44668
diff changeset
  6364
    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
  6365
  hence "continuous_on s g" unfolding continuous_on_iff by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6366
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6367
  hence "((snd \<circ> h \<circ> r) ---> g a) sequentially" unfolding continuous_on_sequentially
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6368
    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
  6369
    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
  6370
  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
  6371
    unfolding `a=b` and o_assoc by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6372
  moreover
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6373
  { fix x assume "x\<in>s" "g x = x" "x\<noteq>a"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6374
    hence "False" using dist[THEN bspec[where x=a], THEN bspec[where x=x]]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6375
      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
  6376
  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
  6377
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6378
44131
5fc334b94e00 declare tendsto_const [intro] (accidentally removed in 230a8665c919)
huffman
parents: 44129
diff changeset
  6379
declare tendsto_const [intro] (* FIXME: move *)
5fc334b94e00 declare tendsto_const [intro] (accidentally removed in 230a8665c919)
huffman
parents: 44129
diff changeset
  6380
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6381
end