src/HOL/Limits.thy
author huffman
Sun, 14 Aug 2011 07:54:24 -0700
changeset 45065 0639898074ae
parent 44952 730f7cced3a6
child 45066 f5363511b212
permissions -rw-r--r--
generalize lemmas about LIM and LIMSEQ to tendsto
     1 (*  Title       : Limits.thy
     2     Author      : Brian Huffman
     3 *)
     4 
     5 header {* Filters and Limits *}
     6 
     7 theory Limits
     8 imports RealVector
     9 begin
    10 
    11 subsection {* Filters *}
    12 
    13 text {*
    14   This definition also allows non-proper filters.
    15 *}
    16 
    17 locale is_filter =
    18   fixes F :: "('a \<Rightarrow> bool) \<Rightarrow> bool"
    19   assumes True: "F (\<lambda>x. True)"
    20   assumes conj: "F (\<lambda>x. P x) \<Longrightarrow> F (\<lambda>x. Q x) \<Longrightarrow> F (\<lambda>x. P x \<and> Q x)"
    21   assumes mono: "\<forall>x. P x \<longrightarrow> Q x \<Longrightarrow> F (\<lambda>x. P x) \<Longrightarrow> F (\<lambda>x. Q x)"
    22 
    23 typedef (open) 'a filter = "{F :: ('a \<Rightarrow> bool) \<Rightarrow> bool. is_filter F}"
    24 proof
    25   show "(\<lambda>x. True) \<in> ?filter" by (auto intro: is_filter.intro)
    26 qed
    27 
    28 lemma is_filter_Rep_filter: "is_filter (Rep_filter A)"
    29   using Rep_filter [of A] by simp
    30 
    31 lemma Abs_filter_inverse':
    32   assumes "is_filter F" shows "Rep_filter (Abs_filter F) = F"
    33   using assms by (simp add: Abs_filter_inverse)
    34 
    35 
    36 subsection {* Eventually *}
    37 
    38 definition eventually :: "('a \<Rightarrow> bool) \<Rightarrow> 'a filter \<Rightarrow> bool"
    39   where "eventually P A \<longleftrightarrow> Rep_filter A P"
    40 
    41 lemma eventually_Abs_filter:
    42   assumes "is_filter F" shows "eventually P (Abs_filter F) = F P"
    43   unfolding eventually_def using assms by (simp add: Abs_filter_inverse)
    44 
    45 lemma filter_eq_iff:
    46   shows "A = B \<longleftrightarrow> (\<forall>P. eventually P A = eventually P B)"
    47   unfolding Rep_filter_inject [symmetric] fun_eq_iff eventually_def ..
    48 
    49 lemma eventually_True [simp]: "eventually (\<lambda>x. True) A"
    50   unfolding eventually_def
    51   by (rule is_filter.True [OF is_filter_Rep_filter])
    52 
    53 lemma always_eventually: "\<forall>x. P x \<Longrightarrow> eventually P A"
    54 proof -
    55   assume "\<forall>x. P x" hence "P = (\<lambda>x. True)" by (simp add: ext)
    56   thus "eventually P A" by simp
    57 qed
    58 
    59 lemma eventually_mono:
    60   "(\<forall>x. P x \<longrightarrow> Q x) \<Longrightarrow> eventually P A \<Longrightarrow> eventually Q A"
    61   unfolding eventually_def
    62   by (rule is_filter.mono [OF is_filter_Rep_filter])
    63 
    64 lemma eventually_conj:
    65   assumes P: "eventually (\<lambda>x. P x) A"
    66   assumes Q: "eventually (\<lambda>x. Q x) A"
    67   shows "eventually (\<lambda>x. P x \<and> Q x) A"
    68   using assms unfolding eventually_def
    69   by (rule is_filter.conj [OF is_filter_Rep_filter])
    70 
    71 lemma eventually_mp:
    72   assumes "eventually (\<lambda>x. P x \<longrightarrow> Q x) A"
    73   assumes "eventually (\<lambda>x. P x) A"
    74   shows "eventually (\<lambda>x. Q x) A"
    75 proof (rule eventually_mono)
    76   show "\<forall>x. (P x \<longrightarrow> Q x) \<and> P x \<longrightarrow> Q x" by simp
    77   show "eventually (\<lambda>x. (P x \<longrightarrow> Q x) \<and> P x) A"
    78     using assms by (rule eventually_conj)
    79 qed
    80 
    81 lemma eventually_rev_mp:
    82   assumes "eventually (\<lambda>x. P x) A"
    83   assumes "eventually (\<lambda>x. P x \<longrightarrow> Q x) A"
    84   shows "eventually (\<lambda>x. Q x) A"
    85 using assms(2) assms(1) by (rule eventually_mp)
    86 
    87 lemma eventually_conj_iff:
    88   "eventually (\<lambda>x. P x \<and> Q x) A \<longleftrightarrow> eventually P A \<and> eventually Q A"
    89   by (auto intro: eventually_conj elim: eventually_rev_mp)
    90 
    91 lemma eventually_elim1:
    92   assumes "eventually (\<lambda>i. P i) A"
    93   assumes "\<And>i. P i \<Longrightarrow> Q i"
    94   shows "eventually (\<lambda>i. Q i) A"
    95   using assms by (auto elim!: eventually_rev_mp)
    96 
    97 lemma eventually_elim2:
    98   assumes "eventually (\<lambda>i. P i) A"
    99   assumes "eventually (\<lambda>i. Q i) A"
   100   assumes "\<And>i. P i \<Longrightarrow> Q i \<Longrightarrow> R i"
   101   shows "eventually (\<lambda>i. R i) A"
   102   using assms by (auto elim!: eventually_rev_mp)
   103 
   104 subsection {* Finer-than relation *}
   105 
   106 text {* @{term "A \<le> B"} means that filter @{term A} is finer than
   107 filter @{term B}. *}
   108 
   109 instantiation filter :: (type) complete_lattice
   110 begin
   111 
   112 definition le_filter_def:
   113   "A \<le> B \<longleftrightarrow> (\<forall>P. eventually P B \<longrightarrow> eventually P A)"
   114 
   115 definition
   116   "(A :: 'a filter) < B \<longleftrightarrow> A \<le> B \<and> \<not> B \<le> A"
   117 
   118 definition
   119   "top = Abs_filter (\<lambda>P. \<forall>x. P x)"
   120 
   121 definition
   122   "bot = Abs_filter (\<lambda>P. True)"
   123 
   124 definition
   125   "sup A B = Abs_filter (\<lambda>P. eventually P A \<and> eventually P B)"
   126 
   127 definition
   128   "inf A B = Abs_filter
   129       (\<lambda>P. \<exists>Q R. eventually Q A \<and> eventually R B \<and> (\<forall>x. Q x \<and> R x \<longrightarrow> P x))"
   130 
   131 definition
   132   "Sup S = Abs_filter (\<lambda>P. \<forall>A\<in>S. eventually P A)"
   133 
   134 definition
   135   "Inf S = Sup {A::'a filter. \<forall>B\<in>S. A \<le> B}"
   136 
   137 lemma eventually_top [simp]: "eventually P top \<longleftrightarrow> (\<forall>x. P x)"
   138   unfolding top_filter_def
   139   by (rule eventually_Abs_filter, rule is_filter.intro, auto)
   140 
   141 lemma eventually_bot [simp]: "eventually P bot"
   142   unfolding bot_filter_def
   143   by (subst eventually_Abs_filter, rule is_filter.intro, auto)
   144 
   145 lemma eventually_sup:
   146   "eventually P (sup A B) \<longleftrightarrow> eventually P A \<and> eventually P B"
   147   unfolding sup_filter_def
   148   by (rule eventually_Abs_filter, rule is_filter.intro)
   149      (auto elim!: eventually_rev_mp)
   150 
   151 lemma eventually_inf:
   152   "eventually P (inf A B) \<longleftrightarrow>
   153    (\<exists>Q R. eventually Q A \<and> eventually R B \<and> (\<forall>x. Q x \<and> R x \<longrightarrow> P x))"
   154   unfolding inf_filter_def
   155   apply (rule eventually_Abs_filter, rule is_filter.intro)
   156   apply (fast intro: eventually_True)
   157   apply clarify
   158   apply (intro exI conjI)
   159   apply (erule (1) eventually_conj)
   160   apply (erule (1) eventually_conj)
   161   apply simp
   162   apply auto
   163   done
   164 
   165 lemma eventually_Sup:
   166   "eventually P (Sup S) \<longleftrightarrow> (\<forall>A\<in>S. eventually P A)"
   167   unfolding Sup_filter_def
   168   apply (rule eventually_Abs_filter, rule is_filter.intro)
   169   apply (auto intro: eventually_conj elim!: eventually_rev_mp)
   170   done
   171 
   172 instance proof
   173   fix A B :: "'a filter" show "A < B \<longleftrightarrow> A \<le> B \<and> \<not> B \<le> A"
   174     by (rule less_filter_def)
   175 next
   176   fix A :: "'a filter" show "A \<le> A"
   177     unfolding le_filter_def by simp
   178 next
   179   fix A B C :: "'a filter" assume "A \<le> B" and "B \<le> C" thus "A \<le> C"
   180     unfolding le_filter_def by simp
   181 next
   182   fix A B :: "'a filter" assume "A \<le> B" and "B \<le> A" thus "A = B"
   183     unfolding le_filter_def filter_eq_iff by fast
   184 next
   185   fix A :: "'a filter" show "A \<le> top"
   186     unfolding le_filter_def eventually_top by (simp add: always_eventually)
   187 next
   188   fix A :: "'a filter" show "bot \<le> A"
   189     unfolding le_filter_def by simp
   190 next
   191   fix A B :: "'a filter" show "A \<le> sup A B" and "B \<le> sup A B"
   192     unfolding le_filter_def eventually_sup by simp_all
   193 next
   194   fix A B C :: "'a filter" assume "A \<le> C" and "B \<le> C" thus "sup A B \<le> C"
   195     unfolding le_filter_def eventually_sup by simp
   196 next
   197   fix A B :: "'a filter" show "inf A B \<le> A" and "inf A B \<le> B"
   198     unfolding le_filter_def eventually_inf by (auto intro: eventually_True)
   199 next
   200   fix A B C :: "'a filter" assume "A \<le> B" and "A \<le> C" thus "A \<le> inf B C"
   201     unfolding le_filter_def eventually_inf
   202     by (auto elim!: eventually_mono intro: eventually_conj)
   203 next
   204   fix A :: "'a filter" and S assume "A \<in> S" thus "A \<le> Sup S"
   205     unfolding le_filter_def eventually_Sup by simp
   206 next
   207   fix S and B :: "'a filter" assume "\<And>A. A \<in> S \<Longrightarrow> A \<le> B" thus "Sup S \<le> B"
   208     unfolding le_filter_def eventually_Sup by simp
   209 next
   210   fix C :: "'a filter" and S assume "C \<in> S" thus "Inf S \<le> C"
   211     unfolding le_filter_def Inf_filter_def eventually_Sup Ball_def by simp
   212 next
   213   fix S and A :: "'a filter" assume "\<And>B. B \<in> S \<Longrightarrow> A \<le> B" thus "A \<le> Inf S"
   214     unfolding le_filter_def Inf_filter_def eventually_Sup Ball_def by simp
   215 qed
   216 
   217 end
   218 
   219 lemma filter_leD:
   220   "A \<le> B \<Longrightarrow> eventually P B \<Longrightarrow> eventually P A"
   221   unfolding le_filter_def by simp
   222 
   223 lemma filter_leI:
   224   "(\<And>P. eventually P B \<Longrightarrow> eventually P A) \<Longrightarrow> A \<le> B"
   225   unfolding le_filter_def by simp
   226 
   227 lemma eventually_False:
   228   "eventually (\<lambda>x. False) A \<longleftrightarrow> A = bot"
   229   unfolding filter_eq_iff by (auto elim: eventually_rev_mp)
   230 
   231 subsection {* Map function for filters *}
   232 
   233 definition filtermap :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a filter \<Rightarrow> 'b filter"
   234   where "filtermap f A = Abs_filter (\<lambda>P. eventually (\<lambda>x. P (f x)) A)"
   235 
   236 lemma eventually_filtermap:
   237   "eventually P (filtermap f A) = eventually (\<lambda>x. P (f x)) A"
   238   unfolding filtermap_def
   239   apply (rule eventually_Abs_filter)
   240   apply (rule is_filter.intro)
   241   apply (auto elim!: eventually_rev_mp)
   242   done
   243 
   244 lemma filtermap_ident: "filtermap (\<lambda>x. x) A = A"
   245   by (simp add: filter_eq_iff eventually_filtermap)
   246 
   247 lemma filtermap_filtermap:
   248   "filtermap f (filtermap g A) = filtermap (\<lambda>x. f (g x)) A"
   249   by (simp add: filter_eq_iff eventually_filtermap)
   250 
   251 lemma filtermap_mono: "A \<le> B \<Longrightarrow> filtermap f A \<le> filtermap f B"
   252   unfolding le_filter_def eventually_filtermap by simp
   253 
   254 lemma filtermap_bot [simp]: "filtermap f bot = bot"
   255   by (simp add: filter_eq_iff eventually_filtermap)
   256 
   257 
   258 subsection {* Sequentially *}
   259 
   260 definition sequentially :: "nat filter"
   261   where "sequentially = Abs_filter (\<lambda>P. \<exists>k. \<forall>n\<ge>k. P n)"
   262 
   263 lemma eventually_sequentially:
   264   "eventually P sequentially \<longleftrightarrow> (\<exists>N. \<forall>n\<ge>N. P n)"
   265 unfolding sequentially_def
   266 proof (rule eventually_Abs_filter, rule is_filter.intro)
   267   fix P Q :: "nat \<Rightarrow> bool"
   268   assume "\<exists>i. \<forall>n\<ge>i. P n" and "\<exists>j. \<forall>n\<ge>j. Q n"
   269   then obtain i j where "\<forall>n\<ge>i. P n" and "\<forall>n\<ge>j. Q n" by auto
   270   then have "\<forall>n\<ge>max i j. P n \<and> Q n" by simp
   271   then show "\<exists>k. \<forall>n\<ge>k. P n \<and> Q n" ..
   272 qed auto
   273 
   274 lemma sequentially_bot [simp]: "sequentially \<noteq> bot"
   275   unfolding filter_eq_iff eventually_sequentially by auto
   276 
   277 lemma eventually_False_sequentially [simp]:
   278   "\<not> eventually (\<lambda>n. False) sequentially"
   279   by (simp add: eventually_False)
   280 
   281 lemma le_sequentially:
   282   "A \<le> sequentially \<longleftrightarrow> (\<forall>N. eventually (\<lambda>n. N \<le> n) A)"
   283   unfolding le_filter_def eventually_sequentially
   284   by (safe, fast, drule_tac x=N in spec, auto elim: eventually_rev_mp)
   285 
   286 
   287 definition trivial_limit :: "'a filter \<Rightarrow> bool"
   288   where "trivial_limit A \<longleftrightarrow> eventually (\<lambda>x. False) A"
   289 
   290 lemma trivial_limit_sequentially [intro]: "\<not> trivial_limit sequentially"
   291   by (auto simp add: trivial_limit_def eventually_sequentially)
   292 
   293 subsection {* Standard filters *}
   294 
   295 definition within :: "'a filter \<Rightarrow> 'a set \<Rightarrow> 'a filter" (infixr "within" 70)
   296   where "A within S = Abs_filter (\<lambda>P. eventually (\<lambda>x. x \<in> S \<longrightarrow> P x) A)"
   297 
   298 definition nhds :: "'a::topological_space \<Rightarrow> 'a filter"
   299   where "nhds a = Abs_filter (\<lambda>P. \<exists>S. open S \<and> a \<in> S \<and> (\<forall>x\<in>S. P x))"
   300 
   301 definition at :: "'a::topological_space \<Rightarrow> 'a filter"
   302   where "at a = nhds a within - {a}"
   303 
   304 lemma eventually_within:
   305   "eventually P (A within S) = eventually (\<lambda>x. x \<in> S \<longrightarrow> P x) A"
   306   unfolding within_def
   307   by (rule eventually_Abs_filter, rule is_filter.intro)
   308      (auto elim!: eventually_rev_mp)
   309 
   310 lemma within_UNIV: "A within UNIV = A"
   311   unfolding filter_eq_iff eventually_within by simp
   312 
   313 lemma eventually_nhds:
   314   "eventually P (nhds a) \<longleftrightarrow> (\<exists>S. open S \<and> a \<in> S \<and> (\<forall>x\<in>S. P x))"
   315 unfolding nhds_def
   316 proof (rule eventually_Abs_filter, rule is_filter.intro)
   317   have "open UNIV \<and> a \<in> UNIV \<and> (\<forall>x\<in>UNIV. True)" by simp
   318   thus "\<exists>S. open S \<and> a \<in> S \<and> (\<forall>x\<in>S. True)" by - rule
   319 next
   320   fix P Q
   321   assume "\<exists>S. open S \<and> a \<in> S \<and> (\<forall>x\<in>S. P x)"
   322      and "\<exists>T. open T \<and> a \<in> T \<and> (\<forall>x\<in>T. Q x)"
   323   then obtain S T where
   324     "open S \<and> a \<in> S \<and> (\<forall>x\<in>S. P x)"
   325     "open T \<and> a \<in> T \<and> (\<forall>x\<in>T. Q x)" by auto
   326   hence "open (S \<inter> T) \<and> a \<in> S \<inter> T \<and> (\<forall>x\<in>(S \<inter> T). P x \<and> Q x)"
   327     by (simp add: open_Int)
   328   thus "\<exists>S. open S \<and> a \<in> S \<and> (\<forall>x\<in>S. P x \<and> Q x)" by - rule
   329 qed auto
   330 
   331 lemma eventually_nhds_metric:
   332   "eventually P (nhds a) \<longleftrightarrow> (\<exists>d>0. \<forall>x. dist x a < d \<longrightarrow> P x)"
   333 unfolding eventually_nhds open_dist
   334 apply safe
   335 apply fast
   336 apply (rule_tac x="{x. dist x a < d}" in exI, simp)
   337 apply clarsimp
   338 apply (rule_tac x="d - dist x a" in exI, clarsimp)
   339 apply (simp only: less_diff_eq)
   340 apply (erule le_less_trans [OF dist_triangle])
   341 done
   342 
   343 lemma eventually_at_topological:
   344   "eventually P (at a) \<longleftrightarrow> (\<exists>S. open S \<and> a \<in> S \<and> (\<forall>x\<in>S. x \<noteq> a \<longrightarrow> P x))"
   345 unfolding at_def eventually_within eventually_nhds by simp
   346 
   347 lemma eventually_at:
   348   fixes a :: "'a::metric_space"
   349   shows "eventually P (at a) \<longleftrightarrow> (\<exists>d>0. \<forall>x. x \<noteq> a \<and> dist x a < d \<longrightarrow> P x)"
   350 unfolding at_def eventually_within eventually_nhds_metric by auto
   351 
   352 
   353 subsection {* Boundedness *}
   354 
   355 definition Bfun :: "('a \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> 'a filter \<Rightarrow> bool"
   356   where "Bfun f A = (\<exists>K>0. eventually (\<lambda>x. norm (f x) \<le> K) A)"
   357 
   358 lemma BfunI:
   359   assumes K: "eventually (\<lambda>x. norm (f x) \<le> K) A" shows "Bfun f A"
   360 unfolding Bfun_def
   361 proof (intro exI conjI allI)
   362   show "0 < max K 1" by simp
   363 next
   364   show "eventually (\<lambda>x. norm (f x) \<le> max K 1) A"
   365     using K by (rule eventually_elim1, simp)
   366 qed
   367 
   368 lemma BfunE:
   369   assumes "Bfun f A"
   370   obtains B where "0 < B" and "eventually (\<lambda>x. norm (f x) \<le> B) A"
   371 using assms unfolding Bfun_def by fast
   372 
   373 
   374 subsection {* Convergence to Zero *}
   375 
   376 definition Zfun :: "('a \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> 'a filter \<Rightarrow> bool"
   377   where "Zfun f A = (\<forall>r>0. eventually (\<lambda>x. norm (f x) < r) A)"
   378 
   379 lemma ZfunI:
   380   "(\<And>r. 0 < r \<Longrightarrow> eventually (\<lambda>x. norm (f x) < r) A) \<Longrightarrow> Zfun f A"
   381   unfolding Zfun_def by simp
   382 
   383 lemma ZfunD:
   384   "\<lbrakk>Zfun f A; 0 < r\<rbrakk> \<Longrightarrow> eventually (\<lambda>x. norm (f x) < r) A"
   385   unfolding Zfun_def by simp
   386 
   387 lemma Zfun_ssubst:
   388   "eventually (\<lambda>x. f x = g x) A \<Longrightarrow> Zfun g A \<Longrightarrow> Zfun f A"
   389   unfolding Zfun_def by (auto elim!: eventually_rev_mp)
   390 
   391 lemma Zfun_zero: "Zfun (\<lambda>x. 0) A"
   392   unfolding Zfun_def by simp
   393 
   394 lemma Zfun_norm_iff: "Zfun (\<lambda>x. norm (f x)) A = Zfun (\<lambda>x. f x) A"
   395   unfolding Zfun_def by simp
   396 
   397 lemma Zfun_imp_Zfun:
   398   assumes f: "Zfun f A"
   399   assumes g: "eventually (\<lambda>x. norm (g x) \<le> norm (f x) * K) A"
   400   shows "Zfun (\<lambda>x. g x) A"
   401 proof (cases)
   402   assume K: "0 < K"
   403   show ?thesis
   404   proof (rule ZfunI)
   405     fix r::real assume "0 < r"
   406     hence "0 < r / K"
   407       using K by (rule divide_pos_pos)
   408     then have "eventually (\<lambda>x. norm (f x) < r / K) A"
   409       using ZfunD [OF f] by fast
   410     with g show "eventually (\<lambda>x. norm (g x) < r) A"
   411     proof (rule eventually_elim2)
   412       fix x
   413       assume *: "norm (g x) \<le> norm (f x) * K"
   414       assume "norm (f x) < r / K"
   415       hence "norm (f x) * K < r"
   416         by (simp add: pos_less_divide_eq K)
   417       thus "norm (g x) < r"
   418         by (simp add: order_le_less_trans [OF *])
   419     qed
   420   qed
   421 next
   422   assume "\<not> 0 < K"
   423   hence K: "K \<le> 0" by (simp only: not_less)
   424   show ?thesis
   425   proof (rule ZfunI)
   426     fix r :: real
   427     assume "0 < r"
   428     from g show "eventually (\<lambda>x. norm (g x) < r) A"
   429     proof (rule eventually_elim1)
   430       fix x
   431       assume "norm (g x) \<le> norm (f x) * K"
   432       also have "\<dots> \<le> norm (f x) * 0"
   433         using K norm_ge_zero by (rule mult_left_mono)
   434       finally show "norm (g x) < r"
   435         using `0 < r` by simp
   436     qed
   437   qed
   438 qed
   439 
   440 lemma Zfun_le: "\<lbrakk>Zfun g A; \<forall>x. norm (f x) \<le> norm (g x)\<rbrakk> \<Longrightarrow> Zfun f A"
   441   by (erule_tac K="1" in Zfun_imp_Zfun, simp)
   442 
   443 lemma Zfun_add:
   444   assumes f: "Zfun f A" and g: "Zfun g A"
   445   shows "Zfun (\<lambda>x. f x + g x) A"
   446 proof (rule ZfunI)
   447   fix r::real assume "0 < r"
   448   hence r: "0 < r / 2" by simp
   449   have "eventually (\<lambda>x. norm (f x) < r/2) A"
   450     using f r by (rule ZfunD)
   451   moreover
   452   have "eventually (\<lambda>x. norm (g x) < r/2) A"
   453     using g r by (rule ZfunD)
   454   ultimately
   455   show "eventually (\<lambda>x. norm (f x + g x) < r) A"
   456   proof (rule eventually_elim2)
   457     fix x
   458     assume *: "norm (f x) < r/2" "norm (g x) < r/2"
   459     have "norm (f x + g x) \<le> norm (f x) + norm (g x)"
   460       by (rule norm_triangle_ineq)
   461     also have "\<dots> < r/2 + r/2"
   462       using * by (rule add_strict_mono)
   463     finally show "norm (f x + g x) < r"
   464       by simp
   465   qed
   466 qed
   467 
   468 lemma Zfun_minus: "Zfun f A \<Longrightarrow> Zfun (\<lambda>x. - f x) A"
   469   unfolding Zfun_def by simp
   470 
   471 lemma Zfun_diff: "\<lbrakk>Zfun f A; Zfun g A\<rbrakk> \<Longrightarrow> Zfun (\<lambda>x. f x - g x) A"
   472   by (simp only: diff_minus Zfun_add Zfun_minus)
   473 
   474 lemma (in bounded_linear) Zfun:
   475   assumes g: "Zfun g A"
   476   shows "Zfun (\<lambda>x. f (g x)) A"
   477 proof -
   478   obtain K where "\<And>x. norm (f x) \<le> norm x * K"
   479     using bounded by fast
   480   then have "eventually (\<lambda>x. norm (f (g x)) \<le> norm (g x) * K) A"
   481     by simp
   482   with g show ?thesis
   483     by (rule Zfun_imp_Zfun)
   484 qed
   485 
   486 lemma (in bounded_bilinear) Zfun:
   487   assumes f: "Zfun f A"
   488   assumes g: "Zfun g A"
   489   shows "Zfun (\<lambda>x. f x ** g x) A"
   490 proof (rule ZfunI)
   491   fix r::real assume r: "0 < r"
   492   obtain K where K: "0 < K"
   493     and norm_le: "\<And>x y. norm (x ** y) \<le> norm x * norm y * K"
   494     using pos_bounded by fast
   495   from K have K': "0 < inverse K"
   496     by (rule positive_imp_inverse_positive)
   497   have "eventually (\<lambda>x. norm (f x) < r) A"
   498     using f r by (rule ZfunD)
   499   moreover
   500   have "eventually (\<lambda>x. norm (g x) < inverse K) A"
   501     using g K' by (rule ZfunD)
   502   ultimately
   503   show "eventually (\<lambda>x. norm (f x ** g x) < r) A"
   504   proof (rule eventually_elim2)
   505     fix x
   506     assume *: "norm (f x) < r" "norm (g x) < inverse K"
   507     have "norm (f x ** g x) \<le> norm (f x) * norm (g x) * K"
   508       by (rule norm_le)
   509     also have "norm (f x) * norm (g x) * K < r * inverse K * K"
   510       by (intro mult_strict_right_mono mult_strict_mono' norm_ge_zero * K)
   511     also from K have "r * inverse K * K = r"
   512       by simp
   513     finally show "norm (f x ** g x) < r" .
   514   qed
   515 qed
   516 
   517 lemma (in bounded_bilinear) Zfun_left:
   518   "Zfun f A \<Longrightarrow> Zfun (\<lambda>x. f x ** a) A"
   519   by (rule bounded_linear_left [THEN bounded_linear.Zfun])
   520 
   521 lemma (in bounded_bilinear) Zfun_right:
   522   "Zfun f A \<Longrightarrow> Zfun (\<lambda>x. a ** f x) A"
   523   by (rule bounded_linear_right [THEN bounded_linear.Zfun])
   524 
   525 lemmas Zfun_mult = mult.Zfun
   526 lemmas Zfun_mult_right = mult.Zfun_right
   527 lemmas Zfun_mult_left = mult.Zfun_left
   528 
   529 
   530 subsection {* Limits *}
   531 
   532 definition tendsto :: "('a \<Rightarrow> 'b::topological_space) \<Rightarrow> 'b \<Rightarrow> 'a filter \<Rightarrow> bool"
   533     (infixr "--->" 55) where
   534   "(f ---> l) A \<longleftrightarrow> (\<forall>S. open S \<longrightarrow> l \<in> S \<longrightarrow> eventually (\<lambda>x. f x \<in> S) A)"
   535 
   536 ML {*
   537 structure Tendsto_Intros = Named_Thms
   538 (
   539   val name = "tendsto_intros"
   540   val description = "introduction rules for tendsto"
   541 )
   542 *}
   543 
   544 setup Tendsto_Intros.setup
   545 
   546 lemma tendsto_mono: "A \<le> A' \<Longrightarrow> (f ---> l) A' \<Longrightarrow> (f ---> l) A"
   547   unfolding tendsto_def le_filter_def by fast
   548 
   549 lemma topological_tendstoI:
   550   "(\<And>S. open S \<Longrightarrow> l \<in> S \<Longrightarrow> eventually (\<lambda>x. f x \<in> S) A)
   551     \<Longrightarrow> (f ---> l) A"
   552   unfolding tendsto_def by auto
   553 
   554 lemma topological_tendstoD:
   555   "(f ---> l) A \<Longrightarrow> open S \<Longrightarrow> l \<in> S \<Longrightarrow> eventually (\<lambda>x. f x \<in> S) A"
   556   unfolding tendsto_def by auto
   557 
   558 lemma tendstoI:
   559   assumes "\<And>e. 0 < e \<Longrightarrow> eventually (\<lambda>x. dist (f x) l < e) A"
   560   shows "(f ---> l) A"
   561   apply (rule topological_tendstoI)
   562   apply (simp add: open_dist)
   563   apply (drule (1) bspec, clarify)
   564   apply (drule assms)
   565   apply (erule eventually_elim1, simp)
   566   done
   567 
   568 lemma tendstoD:
   569   "(f ---> l) A \<Longrightarrow> 0 < e \<Longrightarrow> eventually (\<lambda>x. dist (f x) l < e) A"
   570   apply (drule_tac S="{x. dist x l < e}" in topological_tendstoD)
   571   apply (clarsimp simp add: open_dist)
   572   apply (rule_tac x="e - dist x l" in exI, clarsimp)
   573   apply (simp only: less_diff_eq)
   574   apply (erule le_less_trans [OF dist_triangle])
   575   apply simp
   576   apply simp
   577   done
   578 
   579 lemma tendsto_iff:
   580   "(f ---> l) A \<longleftrightarrow> (\<forall>e>0. eventually (\<lambda>x. dist (f x) l < e) A)"
   581   using tendstoI tendstoD by fast
   582 
   583 lemma tendsto_Zfun_iff: "(f ---> a) A = Zfun (\<lambda>x. f x - a) A"
   584   by (simp only: tendsto_iff Zfun_def dist_norm)
   585 
   586 lemma tendsto_ident_at [tendsto_intros]: "((\<lambda>x. x) ---> a) (at a)"
   587   unfolding tendsto_def eventually_at_topological by auto
   588 
   589 lemma tendsto_ident_at_within [tendsto_intros]:
   590   "((\<lambda>x. x) ---> a) (at a within S)"
   591   unfolding tendsto_def eventually_within eventually_at_topological by auto
   592 
   593 lemma tendsto_const [tendsto_intros]: "((\<lambda>x. k) ---> k) A"
   594   by (simp add: tendsto_def)
   595 
   596 lemma tendsto_const_iff:
   597   fixes k l :: "'a::metric_space"
   598   assumes "A \<noteq> bot" shows "((\<lambda>n. k) ---> l) A \<longleftrightarrow> k = l"
   599   apply (safe intro!: tendsto_const)
   600   apply (rule ccontr)
   601   apply (drule_tac e="dist k l" in tendstoD)
   602   apply (simp add: zero_less_dist_iff)
   603   apply (simp add: eventually_False assms)
   604   done
   605 
   606 lemma tendsto_dist [tendsto_intros]:
   607   assumes f: "(f ---> l) A" and g: "(g ---> m) A"
   608   shows "((\<lambda>x. dist (f x) (g x)) ---> dist l m) A"
   609 proof (rule tendstoI)
   610   fix e :: real assume "0 < e"
   611   hence e2: "0 < e/2" by simp
   612   from tendstoD [OF f e2] tendstoD [OF g e2]
   613   show "eventually (\<lambda>x. dist (dist (f x) (g x)) (dist l m) < e) A"
   614   proof (rule eventually_elim2)
   615     fix x assume "dist (f x) l < e/2" "dist (g x) m < e/2"
   616     then show "dist (dist (f x) (g x)) (dist l m) < e"
   617       unfolding dist_real_def
   618       using dist_triangle2 [of "f x" "g x" "l"]
   619       using dist_triangle2 [of "g x" "l" "m"]
   620       using dist_triangle3 [of "l" "m" "f x"]
   621       using dist_triangle [of "f x" "m" "g x"]
   622       by arith
   623   qed
   624 qed
   625 
   626 subsubsection {* Norms *}
   627 
   628 lemma norm_conv_dist: "norm x = dist x 0"
   629   unfolding dist_norm by simp
   630 
   631 lemma tendsto_norm [tendsto_intros]:
   632   "(f ---> a) A \<Longrightarrow> ((\<lambda>x. norm (f x)) ---> norm a) A"
   633   unfolding norm_conv_dist by (intro tendsto_intros)
   634 
   635 lemma tendsto_norm_zero:
   636   "(f ---> 0) A \<Longrightarrow> ((\<lambda>x. norm (f x)) ---> 0) A"
   637   by (drule tendsto_norm, simp)
   638 
   639 lemma tendsto_norm_zero_cancel:
   640   "((\<lambda>x. norm (f x)) ---> 0) A \<Longrightarrow> (f ---> 0) A"
   641   unfolding tendsto_iff dist_norm by simp
   642 
   643 lemma tendsto_norm_zero_iff:
   644   "((\<lambda>x. norm (f x)) ---> 0) A \<longleftrightarrow> (f ---> 0) A"
   645   unfolding tendsto_iff dist_norm by simp
   646 
   647 lemma tendsto_rabs [tendsto_intros]:
   648   "(f ---> (l::real)) A \<Longrightarrow> ((\<lambda>x. \<bar>f x\<bar>) ---> \<bar>l\<bar>) A"
   649   by (fold real_norm_def, rule tendsto_norm)
   650 
   651 lemma tendsto_rabs_zero:
   652   "(f ---> (0::real)) A \<Longrightarrow> ((\<lambda>x. \<bar>f x\<bar>) ---> 0) A"
   653   by (fold real_norm_def, rule tendsto_norm_zero)
   654 
   655 lemma tendsto_rabs_zero_cancel:
   656   "((\<lambda>x. \<bar>f x\<bar>) ---> (0::real)) A \<Longrightarrow> (f ---> 0) A"
   657   by (fold real_norm_def, rule tendsto_norm_zero_cancel)
   658 
   659 lemma tendsto_rabs_zero_iff:
   660   "((\<lambda>x. \<bar>f x\<bar>) ---> (0::real)) A \<longleftrightarrow> (f ---> 0) A"
   661   by (fold real_norm_def, rule tendsto_norm_zero_iff)
   662 
   663 subsubsection {* Addition and subtraction *}
   664 
   665 lemma tendsto_add [tendsto_intros]:
   666   fixes a b :: "'a::real_normed_vector"
   667   shows "\<lbrakk>(f ---> a) A; (g ---> b) A\<rbrakk> \<Longrightarrow> ((\<lambda>x. f x + g x) ---> a + b) A"
   668   by (simp only: tendsto_Zfun_iff add_diff_add Zfun_add)
   669 
   670 lemma tendsto_add_zero:
   671   fixes f g :: "'a::type \<Rightarrow> 'b::real_normed_vector"
   672   shows "\<lbrakk>(f ---> 0) A; (g ---> 0) A\<rbrakk> \<Longrightarrow> ((\<lambda>x. f x + g x) ---> 0) A"
   673   by (drule (1) tendsto_add, simp)
   674 
   675 lemma tendsto_minus [tendsto_intros]:
   676   fixes a :: "'a::real_normed_vector"
   677   shows "(f ---> a) A \<Longrightarrow> ((\<lambda>x. - f x) ---> - a) A"
   678   by (simp only: tendsto_Zfun_iff minus_diff_minus Zfun_minus)
   679 
   680 lemma tendsto_minus_cancel:
   681   fixes a :: "'a::real_normed_vector"
   682   shows "((\<lambda>x. - f x) ---> - a) A \<Longrightarrow> (f ---> a) A"
   683   by (drule tendsto_minus, simp)
   684 
   685 lemma tendsto_diff [tendsto_intros]:
   686   fixes a b :: "'a::real_normed_vector"
   687   shows "\<lbrakk>(f ---> a) A; (g ---> b) A\<rbrakk> \<Longrightarrow> ((\<lambda>x. f x - g x) ---> a - b) A"
   688   by (simp add: diff_minus tendsto_add tendsto_minus)
   689 
   690 lemma tendsto_setsum [tendsto_intros]:
   691   fixes f :: "'a \<Rightarrow> 'b \<Rightarrow> 'c::real_normed_vector"
   692   assumes "\<And>i. i \<in> S \<Longrightarrow> (f i ---> a i) A"
   693   shows "((\<lambda>x. \<Sum>i\<in>S. f i x) ---> (\<Sum>i\<in>S. a i)) A"
   694 proof (cases "finite S")
   695   assume "finite S" thus ?thesis using assms
   696     by (induct, simp add: tendsto_const, simp add: tendsto_add)
   697 next
   698   assume "\<not> finite S" thus ?thesis
   699     by (simp add: tendsto_const)
   700 qed
   701 
   702 subsubsection {* Linear operators and multiplication *}
   703 
   704 lemma (in bounded_linear) tendsto [tendsto_intros]:
   705   "(g ---> a) A \<Longrightarrow> ((\<lambda>x. f (g x)) ---> f a) A"
   706   by (simp only: tendsto_Zfun_iff diff [symmetric] Zfun)
   707 
   708 lemma (in bounded_linear) tendsto_zero:
   709   "(g ---> 0) A \<Longrightarrow> ((\<lambda>x. f (g x)) ---> 0) A"
   710   by (drule tendsto, simp only: zero)
   711 
   712 lemma (in bounded_bilinear) tendsto [tendsto_intros]:
   713   "\<lbrakk>(f ---> a) A; (g ---> b) A\<rbrakk> \<Longrightarrow> ((\<lambda>x. f x ** g x) ---> a ** b) A"
   714   by (simp only: tendsto_Zfun_iff prod_diff_prod
   715                  Zfun_add Zfun Zfun_left Zfun_right)
   716 
   717 lemma (in bounded_bilinear) tendsto_zero:
   718   assumes f: "(f ---> 0) A"
   719   assumes g: "(g ---> 0) A"
   720   shows "((\<lambda>x. f x ** g x) ---> 0) A"
   721   using tendsto [OF f g] by (simp add: zero_left)
   722 
   723 lemma (in bounded_bilinear) tendsto_left_zero:
   724   "(f ---> 0) A \<Longrightarrow> ((\<lambda>x. f x ** c) ---> 0) A"
   725   by (rule bounded_linear.tendsto_zero [OF bounded_linear_left])
   726 
   727 lemma (in bounded_bilinear) tendsto_right_zero:
   728   "(f ---> 0) A \<Longrightarrow> ((\<lambda>x. c ** f x) ---> 0) A"
   729   by (rule bounded_linear.tendsto_zero [OF bounded_linear_right])
   730 
   731 lemmas tendsto_mult = mult.tendsto
   732 
   733 lemma tendsto_power [tendsto_intros]:
   734   fixes f :: "'a \<Rightarrow> 'b::{power,real_normed_algebra}"
   735   shows "(f ---> a) A \<Longrightarrow> ((\<lambda>x. f x ^ n) ---> a ^ n) A"
   736   by (induct n) (simp_all add: tendsto_const tendsto_mult)
   737 
   738 lemma tendsto_setprod [tendsto_intros]:
   739   fixes f :: "'a \<Rightarrow> 'b \<Rightarrow> 'c::{real_normed_algebra,comm_ring_1}"
   740   assumes "\<And>i. i \<in> S \<Longrightarrow> (f i ---> L i) A"
   741   shows "((\<lambda>x. \<Prod>i\<in>S. f i x) ---> (\<Prod>i\<in>S. L i)) A"
   742 proof (cases "finite S")
   743   assume "finite S" thus ?thesis using assms
   744     by (induct, simp add: tendsto_const, simp add: tendsto_mult)
   745 next
   746   assume "\<not> finite S" thus ?thesis
   747     by (simp add: tendsto_const)
   748 qed
   749 
   750 subsubsection {* Inverse and division *}
   751 
   752 lemma (in bounded_bilinear) Zfun_prod_Bfun:
   753   assumes f: "Zfun f A"
   754   assumes g: "Bfun g A"
   755   shows "Zfun (\<lambda>x. f x ** g x) A"
   756 proof -
   757   obtain K where K: "0 \<le> K"
   758     and norm_le: "\<And>x y. norm (x ** y) \<le> norm x * norm y * K"
   759     using nonneg_bounded by fast
   760   obtain B where B: "0 < B"
   761     and norm_g: "eventually (\<lambda>x. norm (g x) \<le> B) A"
   762     using g by (rule BfunE)
   763   have "eventually (\<lambda>x. norm (f x ** g x) \<le> norm (f x) * (B * K)) A"
   764   using norm_g proof (rule eventually_elim1)
   765     fix x
   766     assume *: "norm (g x) \<le> B"
   767     have "norm (f x ** g x) \<le> norm (f x) * norm (g x) * K"
   768       by (rule norm_le)
   769     also have "\<dots> \<le> norm (f x) * B * K"
   770       by (intro mult_mono' order_refl norm_g norm_ge_zero
   771                 mult_nonneg_nonneg K *)
   772     also have "\<dots> = norm (f x) * (B * K)"
   773       by (rule mult_assoc)
   774     finally show "norm (f x ** g x) \<le> norm (f x) * (B * K)" .
   775   qed
   776   with f show ?thesis
   777     by (rule Zfun_imp_Zfun)
   778 qed
   779 
   780 lemma (in bounded_bilinear) flip:
   781   "bounded_bilinear (\<lambda>x y. y ** x)"
   782   apply default
   783   apply (rule add_right)
   784   apply (rule add_left)
   785   apply (rule scaleR_right)
   786   apply (rule scaleR_left)
   787   apply (subst mult_commute)
   788   using bounded by fast
   789 
   790 lemma (in bounded_bilinear) Bfun_prod_Zfun:
   791   assumes f: "Bfun f A"
   792   assumes g: "Zfun g A"
   793   shows "Zfun (\<lambda>x. f x ** g x) A"
   794   using flip g f by (rule bounded_bilinear.Zfun_prod_Bfun)
   795 
   796 lemma Bfun_inverse_lemma:
   797   fixes x :: "'a::real_normed_div_algebra"
   798   shows "\<lbrakk>r \<le> norm x; 0 < r\<rbrakk> \<Longrightarrow> norm (inverse x) \<le> inverse r"
   799   apply (subst nonzero_norm_inverse, clarsimp)
   800   apply (erule (1) le_imp_inverse_le)
   801   done
   802 
   803 lemma Bfun_inverse:
   804   fixes a :: "'a::real_normed_div_algebra"
   805   assumes f: "(f ---> a) A"
   806   assumes a: "a \<noteq> 0"
   807   shows "Bfun (\<lambda>x. inverse (f x)) A"
   808 proof -
   809   from a have "0 < norm a" by simp
   810   hence "\<exists>r>0. r < norm a" by (rule dense)
   811   then obtain r where r1: "0 < r" and r2: "r < norm a" by fast
   812   have "eventually (\<lambda>x. dist (f x) a < r) A"
   813     using tendstoD [OF f r1] by fast
   814   hence "eventually (\<lambda>x. norm (inverse (f x)) \<le> inverse (norm a - r)) A"
   815   proof (rule eventually_elim1)
   816     fix x
   817     assume "dist (f x) a < r"
   818     hence 1: "norm (f x - a) < r"
   819       by (simp add: dist_norm)
   820     hence 2: "f x \<noteq> 0" using r2 by auto
   821     hence "norm (inverse (f x)) = inverse (norm (f x))"
   822       by (rule nonzero_norm_inverse)
   823     also have "\<dots> \<le> inverse (norm a - r)"
   824     proof (rule le_imp_inverse_le)
   825       show "0 < norm a - r" using r2 by simp
   826     next
   827       have "norm a - norm (f x) \<le> norm (a - f x)"
   828         by (rule norm_triangle_ineq2)
   829       also have "\<dots> = norm (f x - a)"
   830         by (rule norm_minus_commute)
   831       also have "\<dots> < r" using 1 .
   832       finally show "norm a - r \<le> norm (f x)" by simp
   833     qed
   834     finally show "norm (inverse (f x)) \<le> inverse (norm a - r)" .
   835   qed
   836   thus ?thesis by (rule BfunI)
   837 qed
   838 
   839 lemma tendsto_inverse_lemma:
   840   fixes a :: "'a::real_normed_div_algebra"
   841   shows "\<lbrakk>(f ---> a) A; a \<noteq> 0; eventually (\<lambda>x. f x \<noteq> 0) A\<rbrakk>
   842          \<Longrightarrow> ((\<lambda>x. inverse (f x)) ---> inverse a) A"
   843   apply (subst tendsto_Zfun_iff)
   844   apply (rule Zfun_ssubst)
   845   apply (erule eventually_elim1)
   846   apply (erule (1) inverse_diff_inverse)
   847   apply (rule Zfun_minus)
   848   apply (rule Zfun_mult_left)
   849   apply (rule mult.Bfun_prod_Zfun)
   850   apply (erule (1) Bfun_inverse)
   851   apply (simp add: tendsto_Zfun_iff)
   852   done
   853 
   854 lemma tendsto_inverse [tendsto_intros]:
   855   fixes a :: "'a::real_normed_div_algebra"
   856   assumes f: "(f ---> a) A"
   857   assumes a: "a \<noteq> 0"
   858   shows "((\<lambda>x. inverse (f x)) ---> inverse a) A"
   859 proof -
   860   from a have "0 < norm a" by simp
   861   with f have "eventually (\<lambda>x. dist (f x) a < norm a) A"
   862     by (rule tendstoD)
   863   then have "eventually (\<lambda>x. f x \<noteq> 0) A"
   864     unfolding dist_norm by (auto elim!: eventually_elim1)
   865   with f a show ?thesis
   866     by (rule tendsto_inverse_lemma)
   867 qed
   868 
   869 lemma tendsto_divide [tendsto_intros]:
   870   fixes a b :: "'a::real_normed_field"
   871   shows "\<lbrakk>(f ---> a) A; (g ---> b) A; b \<noteq> 0\<rbrakk>
   872     \<Longrightarrow> ((\<lambda>x. f x / g x) ---> a / b) A"
   873   by (simp add: mult.tendsto tendsto_inverse divide_inverse)
   874 
   875 lemma tendsto_sgn [tendsto_intros]:
   876   fixes l :: "'a::real_normed_vector"
   877   shows "\<lbrakk>(f ---> l) A; l \<noteq> 0\<rbrakk> \<Longrightarrow> ((\<lambda>x. sgn (f x)) ---> sgn l) A"
   878   unfolding sgn_div_norm by (simp add: tendsto_intros)
   879 
   880 subsubsection {* Uniqueness *}
   881 
   882 lemma tendsto_unique:
   883   fixes f :: "'a \<Rightarrow> 'b::t2_space"
   884   assumes "\<not> trivial_limit A"  "(f ---> l) A"  "(f ---> l') A"
   885   shows "l = l'"
   886 proof (rule ccontr)
   887   assume "l \<noteq> l'"
   888   obtain U V where "open U" "open V" "l \<in> U" "l' \<in> V" "U \<inter> V = {}"
   889     using hausdorff [OF `l \<noteq> l'`] by fast
   890   have "eventually (\<lambda>x. f x \<in> U) A"
   891     using `(f ---> l) A` `open U` `l \<in> U` by (rule topological_tendstoD)
   892   moreover
   893   have "eventually (\<lambda>x. f x \<in> V) A"
   894     using `(f ---> l') A` `open V` `l' \<in> V` by (rule topological_tendstoD)
   895   ultimately
   896   have "eventually (\<lambda>x. False) A"
   897   proof (rule eventually_elim2)
   898     fix x
   899     assume "f x \<in> U" "f x \<in> V"
   900     hence "f x \<in> U \<inter> V" by simp
   901     with `U \<inter> V = {}` show "False" by simp
   902   qed
   903   with `\<not> trivial_limit A` show "False"
   904     by (simp add: trivial_limit_def)
   905 qed
   906 
   907 end