src/HOL/Analysis/Abstract_Limits.thy
author paulson <lp15@cam.ac.uk>
Thu, 07 Mar 2019 14:08:05 +0000
changeset 69874 11065b70407d
child 69875 03bc14eab432
permissions -rw-r--r--
new material for Analysis
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
69874
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     1
theory Abstract_Limits
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     2
  imports
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     3
    Abstract_Topology
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     4
begin
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     5
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     6
subsection\<open>nhdsin and atin\<close>
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     7
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     8
definition nhdsin :: "'a topology \<Rightarrow> 'a \<Rightarrow> 'a filter"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     9
  where "nhdsin X a =
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    10
           (if a \<in> topspace X then (INF S:{S. openin X S \<and> a \<in> S}. principal S) else bot)"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    11
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    12
definition atin :: "'a topology \<Rightarrow> 'a \<Rightarrow> 'a filter"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    13
  where "atin X a \<equiv> inf (nhdsin X a) (principal (topspace X - {a}))"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    14
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    15
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    16
lemma nhdsin_degenerate [simp]: "a \<notin> topspace X \<Longrightarrow> nhdsin X a = bot"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    17
  and atin_degenerate [simp]: "a \<notin> topspace X \<Longrightarrow> atin X a = bot"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    18
  by (simp_all add: nhdsin_def atin_def)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    19
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    20
lemma eventually_nhdsin:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    21
  "eventually P (nhdsin X a) \<longleftrightarrow> a \<notin> topspace X \<or> (\<exists>S. openin X S \<and> a \<in> S \<and> (\<forall>x\<in>S. P x))"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    22
proof (cases "a \<in> topspace X")
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    23
  case True
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    24
  hence "nhdsin X a = (INF S:{S. openin X S \<and> a \<in> S}. principal S)"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    25
    by (simp add: nhdsin_def)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    26
  also have "eventually P \<dots> \<longleftrightarrow> (\<exists>S. openin X S \<and> a \<in> S \<and> (\<forall>x\<in>S. P x))"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    27
    using True by (subst eventually_INF_base) (auto simp: eventually_principal)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    28
  finally show ?thesis using True by simp
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    29
qed auto
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    30
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    31
lemma eventually_atin:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    32
  "eventually P (atin X a) \<longleftrightarrow> a \<notin> topspace X \<or>
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    33
             (\<exists>U. openin X U \<and> a \<in> U \<and> (\<forall>x \<in> U - {a}. P x))"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    34
proof (cases "a \<in> topspace X")
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    35
  case True
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    36
  hence "eventually P (atin X a) \<longleftrightarrow> (\<exists>S. openin X S \<and>
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    37
           a \<in> S \<and> (\<forall>x\<in>S. x \<in> topspace X \<and> x \<noteq> a \<longrightarrow> P x))"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    38
    by (simp add: atin_def eventually_inf_principal eventually_nhdsin)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    39
  also have "\<dots> \<longleftrightarrow> (\<exists>U. openin X U \<and> a \<in> U \<and> (\<forall>x \<in> U - {a}. P x))"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    40
    using openin_subset by (intro ex_cong) auto
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    41
  finally show ?thesis by (simp add: True)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    42
qed auto
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    43
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    44
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    45
subsection\<open>Limits in a topological space\<close>
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    46
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    47
definition limit :: "'a topology \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'b filter \<Rightarrow> bool" where
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    48
  "limit X f l F \<equiv> l \<in> topspace X \<and> (\<forall>U. openin X U \<and> l \<in> U \<longrightarrow> eventually (\<lambda>x. f x \<in> U) F)"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    49
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    50
lemma limit_euclideanreal_iff [simp]: "limit euclideanreal f l F \<longleftrightarrow> (f \<longlongrightarrow> l) F"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    51
  by (auto simp: limit_def tendsto_def)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    52
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    53
lemma limit_in_topspace: "limit X f l F \<Longrightarrow> l \<in> topspace X"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    54
  by (simp add: limit_def)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    55
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    56
lemma limit_const: "limit X (\<lambda>a. l) l F \<longleftrightarrow> l \<in> topspace X"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    57
  by (simp add: limit_def)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    58
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    59
lemma limit_real_const: "limit euclideanreal (\<lambda>a. l) l F"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    60
  by (simp add: limit_def)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    61
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    62
lemma limit_eventually:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    63
   "\<lbrakk>l \<in> topspace X; eventually (\<lambda>x. f x = l) F\<rbrakk> \<Longrightarrow> limit X f l F"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    64
  by (auto simp: limit_def eventually_mono)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    65
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    66
lemma limit_subsequence:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    67
   "\<lbrakk>strict_mono r; limit X f l sequentially\<rbrakk> \<Longrightarrow> limit X (f \<circ> r) l sequentially"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    68
  unfolding limit_def using eventually_subseq by fastforce
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    69
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    70
lemma limit_subtopology:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    71
  "limit (subtopology X S) f l F
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    72
   \<longleftrightarrow> l \<in> S \<and> eventually (\<lambda>a. f a \<in> S) F \<and> limit X f l F"  (is "?lhs = ?rhs")
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    73
proof (cases "l \<in> S \<inter> topspace X")
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    74
  case True
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    75
  show ?thesis
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    76
  proof
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    77
    assume L: ?lhs
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    78
    with True
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    79
    have "\<forall>\<^sub>F b in F. f b \<in> topspace X \<inter> S"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    80
      by (metis (no_types) limit_def openin_topspace topspace_subtopology)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    81
    with L show ?rhs
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    82
      apply (clarsimp simp add: limit_def eventually_mono topspace_subtopology openin_subtopology_alt)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    83
      apply (drule_tac x="S \<inter> U" in spec, force simp: elim: eventually_mono)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    84
      done
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    85
  next
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    86
    assume ?rhs
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    87
    then show ?lhs
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    88
      using eventually_elim2
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    89
      by (fastforce simp add: limit_def topspace_subtopology openin_subtopology_alt)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    90
  qed
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    91
qed (auto simp: limit_def topspace_subtopology)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    92
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    93
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    94
lemma limit_sequentially:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    95
   "limit X S l sequentially \<longleftrightarrow>
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    96
     l \<in> topspace X \<and> (\<forall>U. openin X U \<and> l \<in> U \<longrightarrow> (\<exists>N. \<forall>n. N \<le> n \<longrightarrow> S n \<in> U))"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    97
  by (simp add: limit_def eventually_sequentially)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    98
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    99
lemma limit_sequentially_offset:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   100
   "limit X f l sequentially \<Longrightarrow> limit X (\<lambda>i. f (i + k)) l sequentially"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   101
  unfolding limit_sequentially
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   102
  by (metis add.commute le_add2 order_trans)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   103
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   104
lemma limit_sequentially_offset_rev:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   105
  assumes "limit X (\<lambda>i. f (i + k)) l sequentially"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   106
  shows "limit X f l sequentially"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   107
proof -
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   108
  have "\<exists>N. \<forall>n\<ge>N. f n \<in> U" if U: "openin X U" "l \<in> U" for U
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   109
  proof -
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   110
    obtain N where "\<And>n. n\<ge>N \<Longrightarrow> f (n + k) \<in> U"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   111
      using assms U unfolding limit_sequentially by blast
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   112
    then have "\<forall>n\<ge>N+k. f n \<in> U"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   113
      by (metis add_leD2 le_add_diff_inverse ordered_cancel_comm_monoid_diff_class.le_diff_conv2 add.commute)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   114
    then show ?thesis ..
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   115
  qed
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   116
  with assms show ?thesis
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   117
    unfolding limit_sequentially
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   118
    by simp
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   119
qed
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   120
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   121
lemma limit_atin:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   122
   "limit Y f y (atin X x) \<longleftrightarrow>
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   123
        y \<in> topspace Y \<and>
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   124
        (x \<in> topspace X
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   125
        \<longrightarrow> (\<forall>V. openin Y V \<and> y \<in> V
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   126
                 \<longrightarrow> (\<exists>U. openin X U \<and> x \<in> U \<and> f ` (U - {x}) \<subseteq> V)))"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   127
  by (auto simp: limit_def eventually_atin image_subset_iff)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   128
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   129
lemma limit_atin_self:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   130
   "limit Y f (f a) (atin X a) \<longleftrightarrow>
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   131
        f a \<in> topspace Y \<and>
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   132
        (a \<in> topspace X
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   133
         \<longrightarrow> (\<forall>V. openin Y V \<and> f a \<in> V
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   134
                  \<longrightarrow> (\<exists>U. openin X U \<and> a \<in> U \<and> f ` U \<subseteq> V)))"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   135
  unfolding limit_atin by fastforce
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   136
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   137
lemma limit_trivial:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   138
   "\<lbrakk>trivial_limit F; y \<in> topspace X\<rbrakk> \<Longrightarrow> limit X f y F"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   139
  by (simp add: limit_def)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   140
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   141
lemma limit_transform_eventually:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   142
   "\<lbrakk>eventually (\<lambda>x. f x = g x) F; limit X f l F\<rbrakk> \<Longrightarrow> limit X g l F"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   143
  unfolding limit_def using eventually_elim2 by fastforce
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   144
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   145
lemma continuous_map_limit:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   146
  assumes "continuous_map X Y g" and f: "limit X f l F"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   147
  shows "limit Y (g \<circ> f) (g l) F"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   148
proof -
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   149
  have "g l \<in> topspace Y"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   150
    by (meson assms continuous_map_def limit_in_topspace)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   151
  moreover
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   152
  have "\<And>U. \<lbrakk>\<forall>V. openin X V \<and> l \<in> V \<longrightarrow> (\<forall>\<^sub>F x in F. f x \<in> V); openin Y U; g l \<in> U\<rbrakk>
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   153
            \<Longrightarrow> \<forall>\<^sub>F x in F. g (f x) \<in> U"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   154
    using assms eventually_mono
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   155
    by (fastforce simp: limit_def dest!: openin_continuous_map_preimage)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   156
  ultimately show ?thesis
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   157
    using f by (fastforce simp add: limit_def)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   158
qed
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   159
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   160
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   161
subsection\<open>Pointwise continuity in topological spaces\<close>
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   162
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   163
definition topcontinuous_at where
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   164
  "topcontinuous_at X Y f x \<longleftrightarrow>
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   165
     x \<in> topspace X \<and>
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   166
     (\<forall>x \<in> topspace X. f x \<in> topspace Y) \<and>
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   167
     (\<forall>V. openin Y V \<and> f x \<in> V
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   168
          \<longrightarrow> (\<exists>U. openin X U \<and> x \<in> U \<and> (\<forall>y \<in> U. f y \<in> V)))"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   169
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   170
lemma topcontinuous_at_atin:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   171
   "topcontinuous_at X Y f x \<longleftrightarrow>
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   172
        x \<in> topspace X \<and>
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   173
        (\<forall>x \<in> topspace X. f x \<in> topspace Y) \<and>
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   174
        limit Y f (f x) (atin X x)"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   175
  unfolding topcontinuous_at_def
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   176
  by (fastforce simp add: limit_atin)+
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   177
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   178
lemma continuous_map_eq_topcontinuous_at:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   179
   "continuous_map X Y f \<longleftrightarrow> (\<forall>x \<in> topspace X. topcontinuous_at X Y f x)"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   180
    (is "?lhs = ?rhs")
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   181
proof
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   182
  assume ?lhs
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   183
  then show ?rhs
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   184
    by (auto simp: continuous_map_def topcontinuous_at_def)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   185
next
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   186
  assume R: ?rhs
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   187
  then show ?lhs
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   188
    apply (auto simp: continuous_map_def topcontinuous_at_def)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   189
    apply (subst openin_subopen, safe)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   190
    apply (drule bspec, assumption)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   191
    using openin_subset[of X] apply (auto simp: subset_iff dest!: spec)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   192
    done
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   193
qed
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   194
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   195
lemma continuous_map_atin:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   196
   "continuous_map X Y f \<longleftrightarrow> (\<forall>x \<in> topspace X. limit Y f (f x) (atin X x))"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   197
  by (auto simp: limit_def topcontinuous_at_atin continuous_map_eq_topcontinuous_at)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   198
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   199
lemma limit_continuous_map:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   200
   "\<lbrakk>continuous_map X Y f; a \<in> topspace X; f a = b\<rbrakk> \<Longrightarrow> limit Y f b (atin X a)"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   201
  by (auto simp: continuous_map_atin)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   202
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   203
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   204
subsection\<open>Combining theorems for continuous functions into the reals\<close>
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   205
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   206
lemma continuous_map_real_const [simp,continuous_intros]:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   207
   "continuous_map X euclideanreal (\<lambda>x. c)"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   208
  by simp
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   209
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   210
lemma continuous_map_real_mult [continuous_intros]:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   211
   "\<lbrakk>continuous_map X euclideanreal f; continuous_map X euclideanreal g\<rbrakk>
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   212
   \<Longrightarrow> continuous_map X euclideanreal (\<lambda>x. f x * g x)"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   213
  by (simp add: continuous_map_atin tendsto_mult)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   214
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   215
lemma continuous_map_real_pow [continuous_intros]:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   216
   "continuous_map X euclideanreal f \<Longrightarrow> continuous_map X euclideanreal (\<lambda>x. f x ^ n)"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   217
  by (induction n) (auto simp: continuous_map_real_mult)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   218
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   219
lemma continuous_map_real_mult_left:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   220
   "continuous_map X euclideanreal f \<Longrightarrow> continuous_map X euclideanreal (\<lambda>x. c * f x)"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   221
  by (simp add: continuous_map_atin tendsto_mult)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   222
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   223
lemma continuous_map_real_mult_left_eq:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   224
   "continuous_map X euclideanreal (\<lambda>x. c * f x) \<longleftrightarrow> c = 0 \<or> continuous_map X euclideanreal f"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   225
proof (cases "c = 0")
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   226
  case False
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   227
  have "continuous_map X euclideanreal (\<lambda>x. c * f x) \<Longrightarrow> continuous_map X euclideanreal f"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   228
    apply (frule continuous_map_real_mult_left [where c="inverse c"])
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   229
    apply (simp add: field_simps False)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   230
    done
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   231
  with False show ?thesis
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   232
    using continuous_map_real_mult_left by blast
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   233
qed simp
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   234
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   235
lemma continuous_map_real_mult_right:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   236
   "continuous_map X euclideanreal f \<Longrightarrow> continuous_map X euclideanreal (\<lambda>x. f x * c)"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   237
  by (simp add: continuous_map_atin tendsto_mult)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   238
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   239
lemma continuous_map_real_mult_right_eq:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   240
   "continuous_map X euclideanreal (\<lambda>x. f x * c) \<longleftrightarrow> c = 0 \<or> continuous_map X euclideanreal f"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   241
  by (simp add: mult.commute flip: continuous_map_real_mult_left_eq)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   242
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   243
lemma continuous_map_real_minus [continuous_intros]:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   244
   "continuous_map X euclideanreal f \<Longrightarrow> continuous_map X euclideanreal (\<lambda>x. - f x)"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   245
  by (simp add: continuous_map_atin tendsto_minus)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   246
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   247
lemma continuous_map_real_minus_eq:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   248
   "continuous_map X euclideanreal (\<lambda>x. - f x) \<longleftrightarrow> continuous_map X euclideanreal f"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   249
  using continuous_map_real_mult_left_eq [where c = "-1"] by auto
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   250
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   251
lemma continuous_map_real_add [continuous_intros]:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   252
   "\<lbrakk>continuous_map X euclideanreal f; continuous_map X euclideanreal g\<rbrakk>
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   253
   \<Longrightarrow> continuous_map X euclideanreal (\<lambda>x. f x + g x)"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   254
  by (simp add: continuous_map_atin tendsto_add)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   255
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   256
lemma continuous_map_real_diff [continuous_intros]:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   257
   "\<lbrakk>continuous_map X euclideanreal f; continuous_map X euclideanreal g\<rbrakk>
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   258
   \<Longrightarrow> continuous_map X euclideanreal (\<lambda>x. f x - g x)"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   259
  by (simp add: continuous_map_atin tendsto_diff)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   260
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   261
lemma continuous_map_real_abs [continuous_intros]:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   262
   "continuous_map X euclideanreal f \<Longrightarrow> continuous_map X euclideanreal (\<lambda>x. abs(f x))"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   263
  by (simp add: continuous_map_atin tendsto_rabs)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   264
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   265
lemma continuous_map_real_max [continuous_intros]:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   266
   "\<lbrakk>continuous_map X euclideanreal f; continuous_map X euclideanreal g\<rbrakk>
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   267
   \<Longrightarrow> continuous_map X euclideanreal (\<lambda>x. max (f x) (g x))"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   268
  by (simp add: continuous_map_atin tendsto_max)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   269
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   270
lemma continuous_map_real_min [continuous_intros]:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   271
   "\<lbrakk>continuous_map X euclideanreal f; continuous_map X euclideanreal g\<rbrakk>
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   272
   \<Longrightarrow> continuous_map X euclideanreal (\<lambda>x. min (f x) (g x))"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   273
  by (simp add: continuous_map_atin tendsto_min)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   274
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   275
lemma continuous_map_sum [continuous_intros]:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   276
   "\<lbrakk>finite I; \<And>i. i \<in> I \<Longrightarrow> continuous_map X euclideanreal (\<lambda>x. f x i)\<rbrakk>
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   277
        \<Longrightarrow> continuous_map X euclideanreal (\<lambda>x. sum (f x) I)"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   278
  by (simp add: continuous_map_atin tendsto_sum)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   279
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   280
lemma continuous_map_prod [continuous_intros]:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   281
   "\<lbrakk>finite I;
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   282
         \<And>i. i \<in> I \<Longrightarrow> continuous_map X euclideanreal (\<lambda>x. f x i)\<rbrakk>
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   283
        \<Longrightarrow> continuous_map X euclideanreal (\<lambda>x. prod (f x) I)"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   284
  by (simp add: continuous_map_atin tendsto_prod)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   285
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   286
lemma continuous_map_real_inverse [continuous_intros]:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   287
   "\<lbrakk>continuous_map X euclideanreal f; \<And>x. x \<in> topspace X \<Longrightarrow> f x \<noteq> 0\<rbrakk>
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   288
        \<Longrightarrow> continuous_map X euclideanreal (\<lambda>x. inverse(f x))"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   289
  by (simp add: continuous_map_atin tendsto_inverse)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   290
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   291
lemma continuous_map_real_divide [continuous_intros]:
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   292
   "\<lbrakk>continuous_map X euclideanreal f; continuous_map X euclideanreal g; \<And>x. x \<in> topspace X \<Longrightarrow> g x \<noteq> 0\<rbrakk>
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   293
   \<Longrightarrow> continuous_map X euclideanreal (\<lambda>x. f x / g x)"
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   294
  by (simp add: continuous_map_atin tendsto_divide)
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   295
11065b70407d new material for Analysis
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   296
end