# HG changeset patch # User huffman # Date 1313336738 25200 # Node ID f5363511b2129fd40b1a4e21f59f2b82c267e752 # Parent 0639898074ae54b3c683ad589dfd5ed5af239715 consistently use variable name 'F' for filters diff -r 0639898074ae -r f5363511b212 src/HOL/Limits.thy --- a/src/HOL/Limits.thy Sun Aug 14 07:54:24 2011 -0700 +++ b/src/HOL/Limits.thy Sun Aug 14 08:45:38 2011 -0700 @@ -25,8 +25,8 @@ show "(\x. True) \ ?filter" by (auto intro: is_filter.intro) qed -lemma is_filter_Rep_filter: "is_filter (Rep_filter A)" - using Rep_filter [of A] by simp +lemma is_filter_Rep_filter: "is_filter (Rep_filter F)" + using Rep_filter [of F] by simp lemma Abs_filter_inverse': assumes "is_filter F" shows "Rep_filter (Abs_filter F) = F" @@ -36,84 +36,84 @@ subsection {* Eventually *} definition eventually :: "('a \ bool) \ 'a filter \ bool" - where "eventually P A \ Rep_filter A P" + where "eventually P F \ Rep_filter F P" lemma eventually_Abs_filter: assumes "is_filter F" shows "eventually P (Abs_filter F) = F P" unfolding eventually_def using assms by (simp add: Abs_filter_inverse) lemma filter_eq_iff: - shows "A = B \ (\P. eventually P A = eventually P B)" + shows "F = F' \ (\P. eventually P F = eventually P F')" unfolding Rep_filter_inject [symmetric] fun_eq_iff eventually_def .. -lemma eventually_True [simp]: "eventually (\x. True) A" +lemma eventually_True [simp]: "eventually (\x. True) F" unfolding eventually_def by (rule is_filter.True [OF is_filter_Rep_filter]) -lemma always_eventually: "\x. P x \ eventually P A" +lemma always_eventually: "\x. P x \ eventually P F" proof - assume "\x. P x" hence "P = (\x. True)" by (simp add: ext) - thus "eventually P A" by simp + thus "eventually P F" by simp qed lemma eventually_mono: - "(\x. P x \ Q x) \ eventually P A \ eventually Q A" + "(\x. P x \ Q x) \ eventually P F \ eventually Q F" unfolding eventually_def by (rule is_filter.mono [OF is_filter_Rep_filter]) lemma eventually_conj: - assumes P: "eventually (\x. P x) A" - assumes Q: "eventually (\x. Q x) A" - shows "eventually (\x. P x \ Q x) A" + assumes P: "eventually (\x. P x) F" + assumes Q: "eventually (\x. Q x) F" + shows "eventually (\x. P x \ Q x) F" using assms unfolding eventually_def by (rule is_filter.conj [OF is_filter_Rep_filter]) lemma eventually_mp: - assumes "eventually (\x. P x \ Q x) A" - assumes "eventually (\x. P x) A" - shows "eventually (\x. Q x) A" + assumes "eventually (\x. P x \ Q x) F" + assumes "eventually (\x. P x) F" + shows "eventually (\x. Q x) F" proof (rule eventually_mono) show "\x. (P x \ Q x) \ P x \ Q x" by simp - show "eventually (\x. (P x \ Q x) \ P x) A" + show "eventually (\x. (P x \ Q x) \ P x) F" using assms by (rule eventually_conj) qed lemma eventually_rev_mp: - assumes "eventually (\x. P x) A" - assumes "eventually (\x. P x \ Q x) A" - shows "eventually (\x. Q x) A" + assumes "eventually (\x. P x) F" + assumes "eventually (\x. P x \ Q x) F" + shows "eventually (\x. Q x) F" using assms(2) assms(1) by (rule eventually_mp) lemma eventually_conj_iff: - "eventually (\x. P x \ Q x) A \ eventually P A \ eventually Q A" + "eventually (\x. P x \ Q x) F \ eventually P F \ eventually Q F" by (auto intro: eventually_conj elim: eventually_rev_mp) lemma eventually_elim1: - assumes "eventually (\i. P i) A" + assumes "eventually (\i. P i) F" assumes "\i. P i \ Q i" - shows "eventually (\i. Q i) A" + shows "eventually (\i. Q i) F" using assms by (auto elim!: eventually_rev_mp) lemma eventually_elim2: - assumes "eventually (\i. P i) A" - assumes "eventually (\i. Q i) A" + assumes "eventually (\i. P i) F" + assumes "eventually (\i. Q i) F" assumes "\i. P i \ Q i \ R i" - shows "eventually (\i. R i) A" + shows "eventually (\i. R i) F" using assms by (auto elim!: eventually_rev_mp) subsection {* Finer-than relation *} -text {* @{term "A \ B"} means that filter @{term A} is finer than -filter @{term B}. *} +text {* @{term "F \ F'"} means that filter @{term F} is finer than +filter @{term F'}. *} instantiation filter :: (type) complete_lattice begin definition le_filter_def: - "A \ B \ (\P. eventually P B \ eventually P A)" + "F \ F' \ (\P. eventually P F' \ eventually P F)" definition - "(A :: 'a filter) < B \ A \ B \ \ B \ A" + "(F :: 'a filter) < F' \ F \ F' \ \ F' \ F" definition "top = Abs_filter (\P. \x. P x)" @@ -122,17 +122,17 @@ "bot = Abs_filter (\P. True)" definition - "sup A B = Abs_filter (\P. eventually P A \ eventually P B)" + "sup F F' = Abs_filter (\P. eventually P F \ eventually P F')" definition - "inf A B = Abs_filter - (\P. \Q R. eventually Q A \ eventually R B \ (\x. Q x \ R x \ P x))" + "inf F F' = Abs_filter + (\P. \Q R. eventually Q F \ eventually R F' \ (\x. Q x \ R x \ P x))" definition - "Sup S = Abs_filter (\P. \A\S. eventually P A)" + "Sup S = Abs_filter (\P. \F\S. eventually P F)" definition - "Inf S = Sup {A::'a filter. \B\S. A \ B}" + "Inf S = Sup {F::'a filter. \F'\S. F \ F'}" lemma eventually_top [simp]: "eventually P top \ (\x. P x)" unfolding top_filter_def @@ -143,14 +143,14 @@ by (subst eventually_Abs_filter, rule is_filter.intro, auto) lemma eventually_sup: - "eventually P (sup A B) \ eventually P A \ eventually P B" + "eventually P (sup F F') \ eventually P F \ eventually P F'" unfolding sup_filter_def by (rule eventually_Abs_filter, rule is_filter.intro) (auto elim!: eventually_rev_mp) lemma eventually_inf: - "eventually P (inf A B) \ - (\Q R. eventually Q A \ eventually R B \ (\x. Q x \ R x \ P x))" + "eventually P (inf F F') \ + (\Q R. eventually Q F \ eventually R F' \ (\x. Q x \ R x \ P x))" unfolding inf_filter_def apply (rule eventually_Abs_filter, rule is_filter.intro) apply (fast intro: eventually_True) @@ -163,92 +163,80 @@ done lemma eventually_Sup: - "eventually P (Sup S) \ (\A\S. eventually P A)" + "eventually P (Sup S) \ (\F\S. eventually P F)" unfolding Sup_filter_def apply (rule eventually_Abs_filter, rule is_filter.intro) apply (auto intro: eventually_conj elim!: eventually_rev_mp) done instance proof - fix A B :: "'a filter" show "A < B \ A \ B \ \ B \ A" - by (rule less_filter_def) -next - fix A :: "'a filter" show "A \ A" - unfolding le_filter_def by simp -next - fix A B C :: "'a filter" assume "A \ B" and "B \ C" thus "A \ C" - unfolding le_filter_def by simp -next - fix A B :: "'a filter" assume "A \ B" and "B \ A" thus "A = B" - unfolding le_filter_def filter_eq_iff by fast -next - fix A :: "'a filter" show "A \ top" - unfolding le_filter_def eventually_top by (simp add: always_eventually) -next - fix A :: "'a filter" show "bot \ A" - unfolding le_filter_def by simp -next - fix A B :: "'a filter" show "A \ sup A B" and "B \ sup A B" - unfolding le_filter_def eventually_sup by simp_all -next - fix A B C :: "'a filter" assume "A \ C" and "B \ C" thus "sup A B \ C" - unfolding le_filter_def eventually_sup by simp -next - fix A B :: "'a filter" show "inf A B \ A" and "inf A B \ B" - unfolding le_filter_def eventually_inf by (auto intro: eventually_True) -next - fix A B C :: "'a filter" assume "A \ B" and "A \ C" thus "A \ inf B C" + fix F F' F'' :: "'a filter" and S :: "'a filter set" + { show "F < F' \ F \ F' \ \ F' \ F" + by (rule less_filter_def) } + { show "F \ F" + unfolding le_filter_def by simp } + { assume "F \ F'" and "F' \ F''" thus "F \ F''" + unfolding le_filter_def by simp } + { assume "F \ F'" and "F' \ F" thus "F = F'" + unfolding le_filter_def filter_eq_iff by fast } + { show "F \ top" + unfolding le_filter_def eventually_top by (simp add: always_eventually) } + { show "bot \ F" + unfolding le_filter_def by simp } + { show "F \ sup F F'" and "F' \ sup F F'" + unfolding le_filter_def eventually_sup by simp_all } + { assume "F \ F''" and "F' \ F''" thus "sup F F' \ F''" + unfolding le_filter_def eventually_sup by simp } + { show "inf F F' \ F" and "inf F F' \ F'" + unfolding le_filter_def eventually_inf by (auto intro: eventually_True) } + { assume "F \ F'" and "F \ F''" thus "F \ inf F' F''" unfolding le_filter_def eventually_inf - by (auto elim!: eventually_mono intro: eventually_conj) -next - fix A :: "'a filter" and S assume "A \ S" thus "A \ Sup S" - unfolding le_filter_def eventually_Sup by simp -next - fix S and B :: "'a filter" assume "\A. A \ S \ A \ B" thus "Sup S \ B" - unfolding le_filter_def eventually_Sup by simp -next - fix C :: "'a filter" and S assume "C \ S" thus "Inf S \ C" - unfolding le_filter_def Inf_filter_def eventually_Sup Ball_def by simp -next - fix S and A :: "'a filter" assume "\B. B \ S \ A \ B" thus "A \ Inf S" - unfolding le_filter_def Inf_filter_def eventually_Sup Ball_def by simp + by (auto elim!: eventually_mono intro: eventually_conj) } + { assume "F \ S" thus "F \ Sup S" + unfolding le_filter_def eventually_Sup by simp } + { assume "\F. F \ S \ F \ F'" thus "Sup S \ F'" + unfolding le_filter_def eventually_Sup by simp } + { assume "F'' \ S" thus "Inf S \ F''" + unfolding le_filter_def Inf_filter_def eventually_Sup Ball_def by simp } + { assume "\F'. F' \ S \ F \ F'" thus "F \ Inf S" + unfolding le_filter_def Inf_filter_def eventually_Sup Ball_def by simp } qed end lemma filter_leD: - "A \ B \ eventually P B \ eventually P A" + "F \ F' \ eventually P F' \ eventually P F" unfolding le_filter_def by simp lemma filter_leI: - "(\P. eventually P B \ eventually P A) \ A \ B" + "(\P. eventually P F' \ eventually P F) \ F \ F'" unfolding le_filter_def by simp lemma eventually_False: - "eventually (\x. False) A \ A = bot" + "eventually (\x. False) F \ F = bot" unfolding filter_eq_iff by (auto elim: eventually_rev_mp) subsection {* Map function for filters *} definition filtermap :: "('a \ 'b) \ 'a filter \ 'b filter" - where "filtermap f A = Abs_filter (\P. eventually (\x. P (f x)) A)" + where "filtermap f F = Abs_filter (\P. eventually (\x. P (f x)) F)" lemma eventually_filtermap: - "eventually P (filtermap f A) = eventually (\x. P (f x)) A" + "eventually P (filtermap f F) = eventually (\x. P (f x)) F" unfolding filtermap_def apply (rule eventually_Abs_filter) apply (rule is_filter.intro) apply (auto elim!: eventually_rev_mp) done -lemma filtermap_ident: "filtermap (\x. x) A = A" +lemma filtermap_ident: "filtermap (\x. x) F = F" by (simp add: filter_eq_iff eventually_filtermap) lemma filtermap_filtermap: - "filtermap f (filtermap g A) = filtermap (\x. f (g x)) A" + "filtermap f (filtermap g F) = filtermap (\x. f (g x)) F" by (simp add: filter_eq_iff eventually_filtermap) -lemma filtermap_mono: "A \ B \ filtermap f A \ filtermap f B" +lemma filtermap_mono: "F \ F' \ filtermap f F \ filtermap f F'" unfolding le_filter_def eventually_filtermap by simp lemma filtermap_bot [simp]: "filtermap f bot = bot" @@ -279,13 +267,13 @@ by (simp add: eventually_False) lemma le_sequentially: - "A \ sequentially \ (\N. eventually (\n. N \ n) A)" + "F \ sequentially \ (\N. eventually (\n. N \ n) F)" unfolding le_filter_def eventually_sequentially by (safe, fast, drule_tac x=N in spec, auto elim: eventually_rev_mp) definition trivial_limit :: "'a filter \ bool" - where "trivial_limit A \ eventually (\x. False) A" + where "trivial_limit F \ eventually (\x. False) F" lemma trivial_limit_sequentially [intro]: "\ trivial_limit sequentially" by (auto simp add: trivial_limit_def eventually_sequentially) @@ -293,7 +281,7 @@ subsection {* Standard filters *} definition within :: "'a filter \ 'a set \ 'a filter" (infixr "within" 70) - where "A within S = Abs_filter (\P. eventually (\x. x \ S \ P x) A)" + where "F within S = Abs_filter (\P. eventually (\x. x \ S \ P x) F)" definition nhds :: "'a::topological_space \ 'a filter" where "nhds a = Abs_filter (\P. \S. open S \ a \ S \ (\x\S. P x))" @@ -302,12 +290,12 @@ where "at a = nhds a within - {a}" lemma eventually_within: - "eventually P (A within S) = eventually (\x. x \ S \ P x) A" + "eventually P (F within S) = eventually (\x. x \ S \ P x) F" unfolding within_def by (rule eventually_Abs_filter, rule is_filter.intro) (auto elim!: eventually_rev_mp) -lemma within_UNIV: "A within UNIV = A" +lemma within_UNIV: "F within UNIV = F" unfolding filter_eq_iff eventually_within by simp lemma eventually_nhds: @@ -353,51 +341,51 @@ subsection {* Boundedness *} definition Bfun :: "('a \ 'b::real_normed_vector) \ 'a filter \ bool" - where "Bfun f A = (\K>0. eventually (\x. norm (f x) \ K) A)" + where "Bfun f F = (\K>0. eventually (\x. norm (f x) \ K) F)" lemma BfunI: - assumes K: "eventually (\x. norm (f x) \ K) A" shows "Bfun f A" + assumes K: "eventually (\x. norm (f x) \ K) F" shows "Bfun f F" unfolding Bfun_def proof (intro exI conjI allI) show "0 < max K 1" by simp next - show "eventually (\x. norm (f x) \ max K 1) A" + show "eventually (\x. norm (f x) \ max K 1) F" using K by (rule eventually_elim1, simp) qed lemma BfunE: - assumes "Bfun f A" - obtains B where "0 < B" and "eventually (\x. norm (f x) \ B) A" + assumes "Bfun f F" + obtains B where "0 < B" and "eventually (\x. norm (f x) \ B) F" using assms unfolding Bfun_def by fast subsection {* Convergence to Zero *} definition Zfun :: "('a \ 'b::real_normed_vector) \ 'a filter \ bool" - where "Zfun f A = (\r>0. eventually (\x. norm (f x) < r) A)" + where "Zfun f F = (\r>0. eventually (\x. norm (f x) < r) F)" lemma ZfunI: - "(\r. 0 < r \ eventually (\x. norm (f x) < r) A) \ Zfun f A" + "(\r. 0 < r \ eventually (\x. norm (f x) < r) F) \ Zfun f F" unfolding Zfun_def by simp lemma ZfunD: - "\Zfun f A; 0 < r\ \ eventually (\x. norm (f x) < r) A" + "\Zfun f F; 0 < r\ \ eventually (\x. norm (f x) < r) F" unfolding Zfun_def by simp lemma Zfun_ssubst: - "eventually (\x. f x = g x) A \ Zfun g A \ Zfun f A" + "eventually (\x. f x = g x) F \ Zfun g F \ Zfun f F" unfolding Zfun_def by (auto elim!: eventually_rev_mp) -lemma Zfun_zero: "Zfun (\x. 0) A" +lemma Zfun_zero: "Zfun (\x. 0) F" unfolding Zfun_def by simp -lemma Zfun_norm_iff: "Zfun (\x. norm (f x)) A = Zfun (\x. f x) A" +lemma Zfun_norm_iff: "Zfun (\x. norm (f x)) F = Zfun (\x. f x) F" unfolding Zfun_def by simp lemma Zfun_imp_Zfun: - assumes f: "Zfun f A" - assumes g: "eventually (\x. norm (g x) \ norm (f x) * K) A" - shows "Zfun (\x. g x) A" + assumes f: "Zfun f F" + assumes g: "eventually (\x. norm (g x) \ norm (f x) * K) F" + shows "Zfun (\x. g x) F" proof (cases) assume K: "0 < K" show ?thesis @@ -405,9 +393,9 @@ fix r::real assume "0 < r" hence "0 < r / K" using K by (rule divide_pos_pos) - then have "eventually (\x. norm (f x) < r / K) A" + then have "eventually (\x. norm (f x) < r / K) F" using ZfunD [OF f] by fast - with g show "eventually (\x. norm (g x) < r) A" + with g show "eventually (\x. norm (g x) < r) F" proof (rule eventually_elim2) fix x assume *: "norm (g x) \ norm (f x) * K" @@ -425,7 +413,7 @@ proof (rule ZfunI) fix r :: real assume "0 < r" - from g show "eventually (\x. norm (g x) < r) A" + from g show "eventually (\x. norm (g x) < r) F" proof (rule eventually_elim1) fix x assume "norm (g x) \ norm (f x) * K" @@ -437,22 +425,22 @@ qed qed -lemma Zfun_le: "\Zfun g A; \x. norm (f x) \ norm (g x)\ \ Zfun f A" +lemma Zfun_le: "\Zfun g F; \x. norm (f x) \ norm (g x)\ \ Zfun f F" by (erule_tac K="1" in Zfun_imp_Zfun, simp) lemma Zfun_add: - assumes f: "Zfun f A" and g: "Zfun g A" - shows "Zfun (\x. f x + g x) A" + assumes f: "Zfun f F" and g: "Zfun g F" + shows "Zfun (\x. f x + g x) F" proof (rule ZfunI) fix r::real assume "0 < r" hence r: "0 < r / 2" by simp - have "eventually (\x. norm (f x) < r/2) A" + have "eventually (\x. norm (f x) < r/2) F" using f r by (rule ZfunD) moreover - have "eventually (\x. norm (g x) < r/2) A" + have "eventually (\x. norm (g x) < r/2) F" using g r by (rule ZfunD) ultimately - show "eventually (\x. norm (f x + g x) < r) A" + show "eventually (\x. norm (f x + g x) < r) F" proof (rule eventually_elim2) fix x assume *: "norm (f x) < r/2" "norm (g x) < r/2" @@ -465,28 +453,28 @@ qed qed -lemma Zfun_minus: "Zfun f A \ Zfun (\x. - f x) A" +lemma Zfun_minus: "Zfun f F \ Zfun (\x. - f x) F" unfolding Zfun_def by simp -lemma Zfun_diff: "\Zfun f A; Zfun g A\ \ Zfun (\x. f x - g x) A" +lemma Zfun_diff: "\Zfun f F; Zfun g F\ \ Zfun (\x. f x - g x) F" by (simp only: diff_minus Zfun_add Zfun_minus) lemma (in bounded_linear) Zfun: - assumes g: "Zfun g A" - shows "Zfun (\x. f (g x)) A" + assumes g: "Zfun g F" + shows "Zfun (\x. f (g x)) F" proof - obtain K where "\x. norm (f x) \ norm x * K" using bounded by fast - then have "eventually (\x. norm (f (g x)) \ norm (g x) * K) A" + then have "eventually (\x. norm (f (g x)) \ norm (g x) * K) F" by simp with g show ?thesis by (rule Zfun_imp_Zfun) qed lemma (in bounded_bilinear) Zfun: - assumes f: "Zfun f A" - assumes g: "Zfun g A" - shows "Zfun (\x. f x ** g x) A" + assumes f: "Zfun f F" + assumes g: "Zfun g F" + shows "Zfun (\x. f x ** g x) F" proof (rule ZfunI) fix r::real assume r: "0 < r" obtain K where K: "0 < K" @@ -494,13 +482,13 @@ using pos_bounded by fast from K have K': "0 < inverse K" by (rule positive_imp_inverse_positive) - have "eventually (\x. norm (f x) < r) A" + have "eventually (\x. norm (f x) < r) F" using f r by (rule ZfunD) moreover - have "eventually (\x. norm (g x) < inverse K) A" + have "eventually (\x. norm (g x) < inverse K) F" using g K' by (rule ZfunD) ultimately - show "eventually (\x. norm (f x ** g x) < r) A" + show "eventually (\x. norm (f x ** g x) < r) F" proof (rule eventually_elim2) fix x assume *: "norm (f x) < r" "norm (g x) < inverse K" @@ -515,11 +503,11 @@ qed lemma (in bounded_bilinear) Zfun_left: - "Zfun f A \ Zfun (\x. f x ** a) A" + "Zfun f F \ Zfun (\x. f x ** a) F" by (rule bounded_linear_left [THEN bounded_linear.Zfun]) lemma (in bounded_bilinear) Zfun_right: - "Zfun f A \ Zfun (\x. a ** f x) A" + "Zfun f F \ Zfun (\x. a ** f x) F" by (rule bounded_linear_right [THEN bounded_linear.Zfun]) lemmas Zfun_mult = mult.Zfun @@ -531,7 +519,7 @@ definition tendsto :: "('a \ 'b::topological_space) \ 'b \ 'a filter \ bool" (infixr "--->" 55) where - "(f ---> l) A \ (\S. open S \ l \ S \ eventually (\x. f x \ S) A)" + "(f ---> l) F \ (\S. open S \ l \ S \ eventually (\x. f x \ S) F)" ML {* structure Tendsto_Intros = Named_Thms @@ -543,21 +531,21 @@ setup Tendsto_Intros.setup -lemma tendsto_mono: "A \ A' \ (f ---> l) A' \ (f ---> l) A" +lemma tendsto_mono: "F \ F' \ (f ---> l) F' \ (f ---> l) F" unfolding tendsto_def le_filter_def by fast lemma topological_tendstoI: - "(\S. open S \ l \ S \ eventually (\x. f x \ S) A) - \ (f ---> l) A" + "(\S. open S \ l \ S \ eventually (\x. f x \ S) F) + \ (f ---> l) F" unfolding tendsto_def by auto lemma topological_tendstoD: - "(f ---> l) A \ open S \ l \ S \ eventually (\x. f x \ S) A" + "(f ---> l) F \ open S \ l \ S \ eventually (\x. f x \ S) F" unfolding tendsto_def by auto lemma tendstoI: - assumes "\e. 0 < e \ eventually (\x. dist (f x) l < e) A" - shows "(f ---> l) A" + assumes "\e. 0 < e \ eventually (\x. dist (f x) l < e) F" + shows "(f ---> l) F" apply (rule topological_tendstoI) apply (simp add: open_dist) apply (drule (1) bspec, clarify) @@ -566,7 +554,7 @@ done lemma tendstoD: - "(f ---> l) A \ 0 < e \ eventually (\x. dist (f x) l < e) A" + "(f ---> l) F \ 0 < e \ eventually (\x. dist (f x) l < e) F" apply (drule_tac S="{x. dist x l < e}" in topological_tendstoD) apply (clarsimp simp add: open_dist) apply (rule_tac x="e - dist x l" in exI, clarsimp) @@ -577,10 +565,10 @@ done lemma tendsto_iff: - "(f ---> l) A \ (\e>0. eventually (\x. dist (f x) l < e) A)" + "(f ---> l) F \ (\e>0. eventually (\x. dist (f x) l < e) F)" using tendstoI tendstoD by fast -lemma tendsto_Zfun_iff: "(f ---> a) A = Zfun (\x. f x - a) A" +lemma tendsto_Zfun_iff: "(f ---> a) F = Zfun (\x. f x - a) F" by (simp only: tendsto_iff Zfun_def dist_norm) lemma tendsto_ident_at [tendsto_intros]: "((\x. x) ---> a) (at a)" @@ -590,12 +578,12 @@ "((\x. x) ---> a) (at a within S)" unfolding tendsto_def eventually_within eventually_at_topological by auto -lemma tendsto_const [tendsto_intros]: "((\x. k) ---> k) A" +lemma tendsto_const [tendsto_intros]: "((\x. k) ---> k) F" by (simp add: tendsto_def) lemma tendsto_const_iff: fixes k l :: "'a::metric_space" - assumes "A \ bot" shows "((\n. k) ---> l) A \ k = l" + assumes "F \ bot" shows "((\n. k) ---> l) F \ k = l" apply (safe intro!: tendsto_const) apply (rule ccontr) apply (drule_tac e="dist k l" in tendstoD) @@ -604,13 +592,13 @@ done lemma tendsto_dist [tendsto_intros]: - assumes f: "(f ---> l) A" and g: "(g ---> m) A" - shows "((\x. dist (f x) (g x)) ---> dist l m) A" + assumes f: "(f ---> l) F" and g: "(g ---> m) F" + shows "((\x. dist (f x) (g x)) ---> dist l m) F" proof (rule tendstoI) fix e :: real assume "0 < e" hence e2: "0 < e/2" by simp from tendstoD [OF f e2] tendstoD [OF g e2] - show "eventually (\x. dist (dist (f x) (g x)) (dist l m) < e) A" + show "eventually (\x. dist (dist (f x) (g x)) (dist l m) < e) F" proof (rule eventually_elim2) fix x assume "dist (f x) l < e/2" "dist (g x) m < e/2" then show "dist (dist (f x) (g x)) (dist l m) < e" @@ -629,68 +617,68 @@ unfolding dist_norm by simp lemma tendsto_norm [tendsto_intros]: - "(f ---> a) A \ ((\x. norm (f x)) ---> norm a) A" + "(f ---> a) F \ ((\x. norm (f x)) ---> norm a) F" unfolding norm_conv_dist by (intro tendsto_intros) lemma tendsto_norm_zero: - "(f ---> 0) A \ ((\x. norm (f x)) ---> 0) A" + "(f ---> 0) F \ ((\x. norm (f x)) ---> 0) F" by (drule tendsto_norm, simp) lemma tendsto_norm_zero_cancel: - "((\x. norm (f x)) ---> 0) A \ (f ---> 0) A" + "((\x. norm (f x)) ---> 0) F \ (f ---> 0) F" unfolding tendsto_iff dist_norm by simp lemma tendsto_norm_zero_iff: - "((\x. norm (f x)) ---> 0) A \ (f ---> 0) A" + "((\x. norm (f x)) ---> 0) F \ (f ---> 0) F" unfolding tendsto_iff dist_norm by simp lemma tendsto_rabs [tendsto_intros]: - "(f ---> (l::real)) A \ ((\x. \f x\) ---> \l\) A" + "(f ---> (l::real)) F \ ((\x. \f x\) ---> \l\) F" by (fold real_norm_def, rule tendsto_norm) lemma tendsto_rabs_zero: - "(f ---> (0::real)) A \ ((\x. \f x\) ---> 0) A" + "(f ---> (0::real)) F \ ((\x. \f x\) ---> 0) F" by (fold real_norm_def, rule tendsto_norm_zero) lemma tendsto_rabs_zero_cancel: - "((\x. \f x\) ---> (0::real)) A \ (f ---> 0) A" + "((\x. \f x\) ---> (0::real)) F \ (f ---> 0) F" by (fold real_norm_def, rule tendsto_norm_zero_cancel) lemma tendsto_rabs_zero_iff: - "((\x. \f x\) ---> (0::real)) A \ (f ---> 0) A" + "((\x. \f x\) ---> (0::real)) F \ (f ---> 0) F" by (fold real_norm_def, rule tendsto_norm_zero_iff) subsubsection {* Addition and subtraction *} lemma tendsto_add [tendsto_intros]: fixes a b :: "'a::real_normed_vector" - shows "\(f ---> a) A; (g ---> b) A\ \ ((\x. f x + g x) ---> a + b) A" + shows "\(f ---> a) F; (g ---> b) F\ \ ((\x. f x + g x) ---> a + b) F" by (simp only: tendsto_Zfun_iff add_diff_add Zfun_add) lemma tendsto_add_zero: fixes f g :: "'a::type \ 'b::real_normed_vector" - shows "\(f ---> 0) A; (g ---> 0) A\ \ ((\x. f x + g x) ---> 0) A" + shows "\(f ---> 0) F; (g ---> 0) F\ \ ((\x. f x + g x) ---> 0) F" by (drule (1) tendsto_add, simp) lemma tendsto_minus [tendsto_intros]: fixes a :: "'a::real_normed_vector" - shows "(f ---> a) A \ ((\x. - f x) ---> - a) A" + shows "(f ---> a) F \ ((\x. - f x) ---> - a) F" by (simp only: tendsto_Zfun_iff minus_diff_minus Zfun_minus) lemma tendsto_minus_cancel: fixes a :: "'a::real_normed_vector" - shows "((\x. - f x) ---> - a) A \ (f ---> a) A" + shows "((\x. - f x) ---> - a) F \ (f ---> a) F" by (drule tendsto_minus, simp) lemma tendsto_diff [tendsto_intros]: fixes a b :: "'a::real_normed_vector" - shows "\(f ---> a) A; (g ---> b) A\ \ ((\x. f x - g x) ---> a - b) A" + shows "\(f ---> a) F; (g ---> b) F\ \ ((\x. f x - g x) ---> a - b) F" by (simp add: diff_minus tendsto_add tendsto_minus) lemma tendsto_setsum [tendsto_intros]: fixes f :: "'a \ 'b \ 'c::real_normed_vector" - assumes "\i. i \ S \ (f i ---> a i) A" - shows "((\x. \i\S. f i x) ---> (\i\S. a i)) A" + assumes "\i. i \ S \ (f i ---> a i) F" + shows "((\x. \i\S. f i x) ---> (\i\S. a i)) F" proof (cases "finite S") assume "finite S" thus ?thesis using assms by (induct, simp add: tendsto_const, simp add: tendsto_add) @@ -702,43 +690,43 @@ subsubsection {* Linear operators and multiplication *} lemma (in bounded_linear) tendsto [tendsto_intros]: - "(g ---> a) A \ ((\x. f (g x)) ---> f a) A" + "(g ---> a) F \ ((\x. f (g x)) ---> f a) F" by (simp only: tendsto_Zfun_iff diff [symmetric] Zfun) lemma (in bounded_linear) tendsto_zero: - "(g ---> 0) A \ ((\x. f (g x)) ---> 0) A" + "(g ---> 0) F \ ((\x. f (g x)) ---> 0) F" by (drule tendsto, simp only: zero) lemma (in bounded_bilinear) tendsto [tendsto_intros]: - "\(f ---> a) A; (g ---> b) A\ \ ((\x. f x ** g x) ---> a ** b) A" + "\(f ---> a) F; (g ---> b) F\ \ ((\x. f x ** g x) ---> a ** b) F" by (simp only: tendsto_Zfun_iff prod_diff_prod Zfun_add Zfun Zfun_left Zfun_right) lemma (in bounded_bilinear) tendsto_zero: - assumes f: "(f ---> 0) A" - assumes g: "(g ---> 0) A" - shows "((\x. f x ** g x) ---> 0) A" + assumes f: "(f ---> 0) F" + assumes g: "(g ---> 0) F" + shows "((\x. f x ** g x) ---> 0) F" using tendsto [OF f g] by (simp add: zero_left) lemma (in bounded_bilinear) tendsto_left_zero: - "(f ---> 0) A \ ((\x. f x ** c) ---> 0) A" + "(f ---> 0) F \ ((\x. f x ** c) ---> 0) F" by (rule bounded_linear.tendsto_zero [OF bounded_linear_left]) lemma (in bounded_bilinear) tendsto_right_zero: - "(f ---> 0) A \ ((\x. c ** f x) ---> 0) A" + "(f ---> 0) F \ ((\x. c ** f x) ---> 0) F" by (rule bounded_linear.tendsto_zero [OF bounded_linear_right]) lemmas tendsto_mult = mult.tendsto lemma tendsto_power [tendsto_intros]: fixes f :: "'a \ 'b::{power,real_normed_algebra}" - shows "(f ---> a) A \ ((\x. f x ^ n) ---> a ^ n) A" + shows "(f ---> a) F \ ((\x. f x ^ n) ---> a ^ n) F" by (induct n) (simp_all add: tendsto_const tendsto_mult) lemma tendsto_setprod [tendsto_intros]: fixes f :: "'a \ 'b \ 'c::{real_normed_algebra,comm_ring_1}" - assumes "\i. i \ S \ (f i ---> L i) A" - shows "((\x. \i\S. f i x) ---> (\i\S. L i)) A" + assumes "\i. i \ S \ (f i ---> L i) F" + shows "((\x. \i\S. f i x) ---> (\i\S. L i)) F" proof (cases "finite S") assume "finite S" thus ?thesis using assms by (induct, simp add: tendsto_const, simp add: tendsto_mult) @@ -750,17 +738,17 @@ subsubsection {* Inverse and division *} lemma (in bounded_bilinear) Zfun_prod_Bfun: - assumes f: "Zfun f A" - assumes g: "Bfun g A" - shows "Zfun (\x. f x ** g x) A" + assumes f: "Zfun f F" + assumes g: "Bfun g F" + shows "Zfun (\x. f x ** g x) F" proof - obtain K where K: "0 \ K" and norm_le: "\x y. norm (x ** y) \ norm x * norm y * K" using nonneg_bounded by fast obtain B where B: "0 < B" - and norm_g: "eventually (\x. norm (g x) \ B) A" + and norm_g: "eventually (\x. norm (g x) \ B) F" using g by (rule BfunE) - have "eventually (\x. norm (f x ** g x) \ norm (f x) * (B * K)) A" + have "eventually (\x. norm (f x ** g x) \ norm (f x) * (B * K)) F" using norm_g proof (rule eventually_elim1) fix x assume *: "norm (g x) \ B" @@ -788,9 +776,9 @@ using bounded by fast lemma (in bounded_bilinear) Bfun_prod_Zfun: - assumes f: "Bfun f A" - assumes g: "Zfun g A" - shows "Zfun (\x. f x ** g x) A" + assumes f: "Bfun f F" + assumes g: "Zfun g F" + shows "Zfun (\x. f x ** g x) F" using flip g f by (rule bounded_bilinear.Zfun_prod_Bfun) lemma Bfun_inverse_lemma: @@ -802,16 +790,16 @@ lemma Bfun_inverse: fixes a :: "'a::real_normed_div_algebra" - assumes f: "(f ---> a) A" + assumes f: "(f ---> a) F" assumes a: "a \ 0" - shows "Bfun (\x. inverse (f x)) A" + shows "Bfun (\x. inverse (f x)) F" proof - from a have "0 < norm a" by simp hence "\r>0. r < norm a" by (rule dense) then obtain r where r1: "0 < r" and r2: "r < norm a" by fast - have "eventually (\x. dist (f x) a < r) A" + have "eventually (\x. dist (f x) a < r) F" using tendstoD [OF f r1] by fast - hence "eventually (\x. norm (inverse (f x)) \ inverse (norm a - r)) A" + hence "eventually (\x. norm (inverse (f x)) \ inverse (norm a - r)) F" proof (rule eventually_elim1) fix x assume "dist (f x) a < r" @@ -838,8 +826,8 @@ lemma tendsto_inverse_lemma: fixes a :: "'a::real_normed_div_algebra" - shows "\(f ---> a) A; a \ 0; eventually (\x. f x \ 0) A\ - \ ((\x. inverse (f x)) ---> inverse a) A" + shows "\(f ---> a) F; a \ 0; eventually (\x. f x \ 0) F\ + \ ((\x. inverse (f x)) ---> inverse a) F" apply (subst tendsto_Zfun_iff) apply (rule Zfun_ssubst) apply (erule eventually_elim1) @@ -853,14 +841,14 @@ lemma tendsto_inverse [tendsto_intros]: fixes a :: "'a::real_normed_div_algebra" - assumes f: "(f ---> a) A" + assumes f: "(f ---> a) F" assumes a: "a \ 0" - shows "((\x. inverse (f x)) ---> inverse a) A" + shows "((\x. inverse (f x)) ---> inverse a) F" proof - from a have "0 < norm a" by simp - with f have "eventually (\x. dist (f x) a < norm a) A" + with f have "eventually (\x. dist (f x) a < norm a) F" by (rule tendstoD) - then have "eventually (\x. f x \ 0) A" + then have "eventually (\x. f x \ 0) F" unfolding dist_norm by (auto elim!: eventually_elim1) with f a show ?thesis by (rule tendsto_inverse_lemma) @@ -868,39 +856,39 @@ lemma tendsto_divide [tendsto_intros]: fixes a b :: "'a::real_normed_field" - shows "\(f ---> a) A; (g ---> b) A; b \ 0\ - \ ((\x. f x / g x) ---> a / b) A" + shows "\(f ---> a) F; (g ---> b) F; b \ 0\ + \ ((\x. f x / g x) ---> a / b) F" by (simp add: mult.tendsto tendsto_inverse divide_inverse) lemma tendsto_sgn [tendsto_intros]: fixes l :: "'a::real_normed_vector" - shows "\(f ---> l) A; l \ 0\ \ ((\x. sgn (f x)) ---> sgn l) A" + shows "\(f ---> l) F; l \ 0\ \ ((\x. sgn (f x)) ---> sgn l) F" unfolding sgn_div_norm by (simp add: tendsto_intros) subsubsection {* Uniqueness *} lemma tendsto_unique: fixes f :: "'a \ 'b::t2_space" - assumes "\ trivial_limit A" "(f ---> l) A" "(f ---> l') A" + assumes "\ trivial_limit F" "(f ---> l) F" "(f ---> l') F" shows "l = l'" proof (rule ccontr) assume "l \ l'" obtain U V where "open U" "open V" "l \ U" "l' \ V" "U \ V = {}" using hausdorff [OF `l \ l'`] by fast - have "eventually (\x. f x \ U) A" - using `(f ---> l) A` `open U` `l \ U` by (rule topological_tendstoD) + have "eventually (\x. f x \ U) F" + using `(f ---> l) F` `open U` `l \ U` by (rule topological_tendstoD) moreover - have "eventually (\x. f x \ V) A" - using `(f ---> l') A` `open V` `l' \ V` by (rule topological_tendstoD) + have "eventually (\x. f x \ V) F" + using `(f ---> l') F` `open V` `l' \ V` by (rule topological_tendstoD) ultimately - have "eventually (\x. False) A" + have "eventually (\x. False) F" proof (rule eventually_elim2) fix x assume "f x \ U" "f x \ V" hence "f x \ U \ V" by simp with `U \ V = {}` show "False" by simp qed - with `\ trivial_limit A` show "False" + with `\ trivial_limit F` show "False" by (simp add: trivial_limit_def) qed