# HG changeset patch # User huffman # Date 1244242760 25200 # Node ID 93938cafc0e61548a02a44e35fad70abe689a13a # Parent 40f815edbde44e01699638669b14b6d92f3da845 put syntax for tendsto in Limits.thy; rename variables diff -r 40f815edbde4 -r 93938cafc0e6 src/HOL/Library/Topology_Euclidean_Space.thy --- a/src/HOL/Library/Topology_Euclidean_Space.thy Fri Jun 05 21:55:08 2009 +0200 +++ b/src/HOL/Library/Topology_Euclidean_Space.thy Fri Jun 05 15:59:20 2009 -0700 @@ -1170,8 +1170,6 @@ subsection{* Limits, defined as vacuously true when the limit is trivial. *} -notation tendsto (infixr "--->" 55) - text{* Notation Lim to avoid collition with lim defined in analysis *} definition "Lim net f = (THE l. (f ---> l) net)" diff -r 40f815edbde4 -r 93938cafc0e6 src/HOL/Lim.thy --- a/src/HOL/Lim.thy Fri Jun 05 21:55:08 2009 +0200 +++ b/src/HOL/Lim.thy Fri Jun 05 15:59:20 2009 -0700 @@ -29,7 +29,7 @@ subsection {* Limits of Functions *} -lemma LIM_conv_tendsto: "(f -- a --> L) \ tendsto f L (at a)" +lemma LIM_conv_tendsto: "(f -- a --> L) \ (f ---> L) (at a)" unfolding LIM_def tendsto_def eventually_at .. lemma metric_LIM_I: diff -r 40f815edbde4 -r 93938cafc0e6 src/HOL/Limits.thy --- a/src/HOL/Limits.thy Fri Jun 05 21:55:08 2009 +0200 +++ b/src/HOL/Limits.thy Fri Jun 05 15:59:20 2009 -0700 @@ -175,20 +175,21 @@ definition Bfun :: "('a \ 'b::real_normed_vector) \ 'a net \ bool" where - [code del]: "Bfun S net = (\K>0. eventually (\i. norm (S i) \ K) net)" + [code del]: "Bfun f net = (\K>0. eventually (\x. norm (f x) \ K) net)" -lemma BfunI: assumes K: "eventually (\i. norm (X i) \ K) net" shows "Bfun X net" +lemma BfunI: + assumes K: "eventually (\x. norm (f x) \ K) net" shows "Bfun f net" unfolding Bfun_def proof (intro exI conjI allI) show "0 < max K 1" by simp next - show "eventually (\i. norm (X i) \ max K 1) net" + show "eventually (\x. norm (f x) \ max K 1) net" using K by (rule eventually_elim1, simp) qed lemma BfunE: - assumes "Bfun S net" - obtains B where "0 < B" and "eventually (\i. norm (S i) \ B) net" + assumes "Bfun f net" + obtains B where "0 < B" and "eventually (\x. norm (f x) \ B) net" using assms unfolding Bfun_def by fast @@ -196,30 +197,30 @@ definition Zfun :: "('a \ 'b::real_normed_vector) \ 'a net \ bool" where - [code del]: "Zfun S net = (\r>0. eventually (\i. norm (S i) < r) net)" + [code del]: "Zfun f net = (\r>0. eventually (\x. norm (f x) < r) net)" lemma ZfunI: - "(\r. 0 < r \ eventually (\i. norm (S i) < r) net) \ Zfun S net" + "(\r. 0 < r \ eventually (\x. norm (f x) < r) net) \ Zfun f net" unfolding Zfun_def by simp lemma ZfunD: - "\Zfun S net; 0 < r\ \ eventually (\i. norm (S i) < r) net" + "\Zfun f net; 0 < r\ \ eventually (\x. norm (f x) < r) net" unfolding Zfun_def by simp lemma Zfun_ssubst: - "eventually (\i. X i = Y i) net \ Zfun Y net \ Zfun X net" + "eventually (\x. f x = g x) net \ Zfun g net \ Zfun f net" unfolding Zfun_def by (auto elim!: eventually_rev_mp) -lemma Zfun_zero: "Zfun (\i. 0) net" +lemma Zfun_zero: "Zfun (\x. 0) net" unfolding Zfun_def by simp -lemma Zfun_norm_iff: "Zfun (\i. norm (S i)) net = Zfun (\i. S i) net" +lemma Zfun_norm_iff: "Zfun (\x. norm (f x)) net = Zfun (\x. f x) net" unfolding Zfun_def by simp lemma Zfun_imp_Zfun: - assumes X: "Zfun X net" - assumes Y: "eventually (\i. norm (Y i) \ norm (X i) * K) net" - shows "Zfun (\n. Y n) net" + assumes f: "Zfun f net" + assumes g: "eventually (\x. norm (g x) \ norm (f x) * K) net" + shows "Zfun (\x. g x) net" proof (cases) assume K: "0 < K" show ?thesis @@ -227,16 +228,16 @@ fix r::real assume "0 < r" hence "0 < r / K" using K by (rule divide_pos_pos) - then have "eventually (\i. norm (X i) < r / K) net" - using ZfunD [OF X] by fast - with Y show "eventually (\i. norm (Y i) < r) net" + then have "eventually (\x. norm (f x) < r / K) net" + using ZfunD [OF f] by fast + with g show "eventually (\x. norm (g x) < r) net" proof (rule eventually_elim2) - fix i - assume *: "norm (Y i) \ norm (X i) * K" - assume "norm (X i) < r / K" - hence "norm (X i) * K < r" + fix x + assume *: "norm (g x) \ norm (f x) * K" + assume "norm (f x) < r / K" + hence "norm (f x) * K < r" by (simp add: pos_less_divide_eq K) - thus "norm (Y i) < r" + thus "norm (g x) < r" by (simp add: order_le_less_trans [OF *]) qed qed @@ -247,68 +248,68 @@ proof (rule ZfunI) fix r :: real assume "0 < r" - from Y show "eventually (\i. norm (Y i) < r) net" + from g show "eventually (\x. norm (g x) < r) net" proof (rule eventually_elim1) - fix i - assume "norm (Y i) \ norm (X i) * K" - also have "\ \ norm (X i) * 0" + fix x + assume "norm (g x) \ norm (f x) * K" + also have "\ \ norm (f x) * 0" using K norm_ge_zero by (rule mult_left_mono) - finally show "norm (Y i) < r" + finally show "norm (g x) < r" using `0 < r` by simp qed qed qed -lemma Zfun_le: "\Zfun Y net; \n. norm (X n) \ norm (Y n)\ \ Zfun X net" +lemma Zfun_le: "\Zfun g net; \x. norm (f x) \ norm (g x)\ \ Zfun f net" by (erule_tac K="1" in Zfun_imp_Zfun, simp) lemma Zfun_add: - assumes X: "Zfun X net" and Y: "Zfun Y net" - shows "Zfun (\n. X n + Y n) net" + assumes f: "Zfun f net" and g: "Zfun g net" + shows "Zfun (\x. f x + g x) net" proof (rule ZfunI) fix r::real assume "0 < r" hence r: "0 < r / 2" by simp - have "eventually (\i. norm (X i) < r/2) net" - using X r by (rule ZfunD) + have "eventually (\x. norm (f x) < r/2) net" + using f r by (rule ZfunD) moreover - have "eventually (\i. norm (Y i) < r/2) net" - using Y r by (rule ZfunD) + have "eventually (\x. norm (g x) < r/2) net" + using g r by (rule ZfunD) ultimately - show "eventually (\i. norm (X i + Y i) < r) net" + show "eventually (\x. norm (f x + g x) < r) net" proof (rule eventually_elim2) - fix i - assume *: "norm (X i) < r/2" "norm (Y i) < r/2" - have "norm (X i + Y i) \ norm (X i) + norm (Y i)" + fix x + assume *: "norm (f x) < r/2" "norm (g x) < r/2" + have "norm (f x + g x) \ norm (f x) + norm (g x)" by (rule norm_triangle_ineq) also have "\ < r/2 + r/2" using * by (rule add_strict_mono) - finally show "norm (X i + Y i) < r" + finally show "norm (f x + g x) < r" by simp qed qed -lemma Zfun_minus: "Zfun X net \ Zfun (\i. - X i) net" +lemma Zfun_minus: "Zfun f net \ Zfun (\x. - f x) net" unfolding Zfun_def by simp -lemma Zfun_diff: "\Zfun X net; Zfun Y net\ \ Zfun (\i. X i - Y i) net" +lemma Zfun_diff: "\Zfun f net; Zfun g net\ \ Zfun (\x. f x - g x) net" by (simp only: diff_minus Zfun_add Zfun_minus) lemma (in bounded_linear) Zfun: - assumes X: "Zfun X net" - shows "Zfun (\n. f (X n)) net" + assumes g: "Zfun g net" + shows "Zfun (\x. f (g x)) net" proof - obtain K where "\x. norm (f x) \ norm x * K" using bounded by fast - then have "eventually (\i. norm (f (X i)) \ norm (X i) * K) net" + then have "eventually (\x. norm (f (g x)) \ norm (g x) * K) net" by simp - with X show ?thesis + with g show ?thesis by (rule Zfun_imp_Zfun) qed lemma (in bounded_bilinear) Zfun: - assumes X: "Zfun X net" - assumes Y: "Zfun Y net" - shows "Zfun (\n. X n ** Y n) net" + assumes f: "Zfun f net" + assumes g: "Zfun g net" + shows "Zfun (\x. f x ** g x) net" proof (rule ZfunI) fix r::real assume r: "0 < r" obtain K where K: "0 < K" @@ -316,32 +317,32 @@ using pos_bounded by fast from K have K': "0 < inverse K" by (rule positive_imp_inverse_positive) - have "eventually (\i. norm (X i) < r) net" - using X r by (rule ZfunD) + have "eventually (\x. norm (f x) < r) net" + using f r by (rule ZfunD) moreover - have "eventually (\i. norm (Y i) < inverse K) net" - using Y K' by (rule ZfunD) + have "eventually (\x. norm (g x) < inverse K) net" + using g K' by (rule ZfunD) ultimately - show "eventually (\i. norm (X i ** Y i) < r) net" + show "eventually (\x. norm (f x ** g x) < r) net" proof (rule eventually_elim2) - fix i - assume *: "norm (X i) < r" "norm (Y i) < inverse K" - have "norm (X i ** Y i) \ norm (X i) * norm (Y i) * K" + fix x + assume *: "norm (f x) < r" "norm (g x) < inverse K" + have "norm (f x ** g x) \ norm (f x) * norm (g x) * K" by (rule norm_le) - also have "norm (X i) * norm (Y i) * K < r * inverse K * K" + also have "norm (f x) * norm (g x) * K < r * inverse K * K" by (intro mult_strict_right_mono mult_strict_mono' norm_ge_zero * K) also from K have "r * inverse K * K = r" by simp - finally show "norm (X i ** Y i) < r" . + finally show "norm (f x ** g x) < r" . qed qed lemma (in bounded_bilinear) Zfun_left: - "Zfun X net \ Zfun (\n. X n ** a) net" + "Zfun f net \ Zfun (\x. f x ** a) net" by (rule bounded_linear_left [THEN bounded_linear.Zfun]) lemma (in bounded_bilinear) Zfun_right: - "Zfun X net \ Zfun (\n. a ** X n) net" + "Zfun f net \ Zfun (\x. a ** f x) net" by (rule bounded_linear_right [THEN bounded_linear.Zfun]) lemmas Zfun_mult = mult.Zfun @@ -352,27 +353,28 @@ subsection{* Limits *} definition - tendsto :: "('a \ 'b::metric_space) \ 'b \ 'a net \ bool" where - [code del]: "tendsto f l net \ (\e>0. eventually (\x. dist (f x) l < e) net)" + tendsto :: "('a \ 'b::metric_space) \ 'b \ 'a net \ bool" + (infixr "--->" 55) where + [code del]: "(f ---> l) net \ (\e>0. eventually (\x. dist (f x) l < e) net)" lemma tendstoI: "(\e. 0 < e \ eventually (\x. dist (f x) l < e) net) - \ tendsto f l net" + \ (f ---> l) net" unfolding tendsto_def by auto lemma tendstoD: - "tendsto f l net \ 0 < e \ eventually (\x. dist (f x) l < e) net" + "(f ---> l) net \ 0 < e \ eventually (\x. dist (f x) l < e) net" unfolding tendsto_def by auto -lemma tendsto_Zfun_iff: "tendsto (\n. X n) L net = Zfun (\n. X n - L) net" +lemma tendsto_Zfun_iff: "(f ---> a) net = Zfun (\x. f x - a) net" by (simp only: tendsto_def Zfun_def dist_norm) -lemma tendsto_const: "tendsto (\n. k) k net" +lemma tendsto_const: "((\x. k) ---> k) net" by (simp add: tendsto_def) lemma tendsto_norm: fixes a :: "'a::real_normed_vector" - shows "tendsto X a net \ tendsto (\n. norm (X n)) (norm a) net" + shows "(f ---> a) net \ ((\x. norm (f x)) ---> norm a) net" apply (simp add: tendsto_def dist_norm, safe) apply (drule_tac x="e" in spec, safe) apply (erule eventually_elim1) @@ -391,30 +393,30 @@ lemma tendsto_add: fixes a b :: "'a::real_normed_vector" - shows "\tendsto X a net; tendsto Y b net\ \ tendsto (\n. X n + Y n) (a + b) net" + shows "\(f ---> a) net; (g ---> b) net\ \ ((\x. f x + g x) ---> a + b) net" by (simp only: tendsto_Zfun_iff add_diff_add Zfun_add) lemma tendsto_minus: fixes a :: "'a::real_normed_vector" - shows "tendsto X a net \ tendsto (\n. - X n) (- a) net" + shows "(f ---> a) net \ ((\x. - f x) ---> - a) net" by (simp only: tendsto_Zfun_iff minus_diff_minus Zfun_minus) lemma tendsto_minus_cancel: fixes a :: "'a::real_normed_vector" - shows "tendsto (\n. - X n) (- a) net \ tendsto X a net" + shows "((\x. - f x) ---> - a) net \ (f ---> a) net" by (drule tendsto_minus, simp) lemma tendsto_diff: fixes a b :: "'a::real_normed_vector" - shows "\tendsto X a net; tendsto Y b net\ \ tendsto (\n. X n - Y n) (a - b) net" + shows "\(f ---> a) net; (g ---> b) net\ \ ((\x. f x - g x) ---> a - b) net" by (simp add: diff_minus tendsto_add tendsto_minus) lemma (in bounded_linear) tendsto: - "tendsto X a net \ tendsto (\n. f (X n)) (f a) net" + "(g ---> a) net \ ((\x. f (g x)) ---> f a) net" by (simp only: tendsto_Zfun_iff diff [symmetric] Zfun) lemma (in bounded_bilinear) tendsto: - "\tendsto X a net; tendsto Y b net\ \ tendsto (\n. X n ** Y n) (a ** b) net" + "\(f ---> a) net; (g ---> b) net\ \ ((\x. f x ** g x) ---> a ** b) net" by (simp only: tendsto_Zfun_iff prod_diff_prod Zfun_add Zfun Zfun_left Zfun_right) @@ -422,31 +424,31 @@ subsection {* Continuity of Inverse *} lemma (in bounded_bilinear) Zfun_prod_Bfun: - assumes X: "Zfun X net" - assumes Y: "Bfun Y net" - shows "Zfun (\n. X n ** Y n) net" + assumes f: "Zfun f net" + assumes g: "Bfun g net" + shows "Zfun (\x. f x ** g x) net" 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_Y: "eventually (\i. norm (Y i) \ B) net" - using Y by (rule BfunE) - have "eventually (\i. norm (X i ** Y i) \ norm (X i) * (B * K)) net" - using norm_Y proof (rule eventually_elim1) - fix i - assume *: "norm (Y i) \ B" - have "norm (X i ** Y i) \ norm (X i) * norm (Y i) * K" + and norm_g: "eventually (\x. norm (g x) \ B) net" + using g by (rule BfunE) + have "eventually (\x. norm (f x ** g x) \ norm (f x) * (B * K)) net" + using norm_g proof (rule eventually_elim1) + fix x + assume *: "norm (g x) \ B" + have "norm (f x ** g x) \ norm (f x) * norm (g x) * K" by (rule norm_le) - also have "\ \ norm (X i) * B * K" - by (intro mult_mono' order_refl norm_Y norm_ge_zero + also have "\ \ norm (f x) * B * K" + by (intro mult_mono' order_refl norm_g norm_ge_zero mult_nonneg_nonneg K *) - also have "\ = norm (X i) * (B * K)" + also have "\ = norm (f x) * (B * K)" by (rule mult_assoc) - finally show "norm (X i ** Y i) \ norm (X i) * (B * K)" . + finally show "norm (f x ** g x) \ norm (f x) * (B * K)" . qed - with X show ?thesis - by (rule Zfun_imp_Zfun) + with f show ?thesis + by (rule Zfun_imp_Zfun) qed lemma (in bounded_bilinear) flip: @@ -460,10 +462,10 @@ using bounded by fast lemma (in bounded_bilinear) Bfun_prod_Zfun: - assumes X: "Bfun X net" - assumes Y: "Zfun Y net" - shows "Zfun (\n. X n ** Y n) net" -using flip Y X by (rule bounded_bilinear.Zfun_prod_Bfun) + assumes f: "Bfun f net" + assumes g: "Zfun g net" + shows "Zfun (\x. f x ** g x) net" +using flip g f by (rule bounded_bilinear.Zfun_prod_Bfun) lemma inverse_diff_inverse: "\(a::'a::division_ring) \ 0; b \ 0\ @@ -479,44 +481,44 @@ lemma Bfun_inverse: fixes a :: "'a::real_normed_div_algebra" - assumes X: "tendsto X a net" + assumes f: "(f ---> a) net" assumes a: "a \ 0" - shows "Bfun (\n. inverse (X n)) net" + shows "Bfun (\x. inverse (f x)) net" 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 (\i. dist (X i) a < r) net" - using tendstoD [OF X r1] by fast - hence "eventually (\i. norm (inverse (X i)) \ inverse (norm a - r)) net" + have "eventually (\x. dist (f x) a < r) net" + using tendstoD [OF f r1] by fast + hence "eventually (\x. norm (inverse (f x)) \ inverse (norm a - r)) net" proof (rule eventually_elim1) - fix i - assume "dist (X i) a < r" - hence 1: "norm (X i - a) < r" + fix x + assume "dist (f x) a < r" + hence 1: "norm (f x - a) < r" by (simp add: dist_norm) - hence 2: "X i \ 0" using r2 by auto - hence "norm (inverse (X i)) = inverse (norm (X i))" + hence 2: "f x \ 0" using r2 by auto + hence "norm (inverse (f x)) = inverse (norm (f x))" by (rule nonzero_norm_inverse) also have "\ \ inverse (norm a - r)" proof (rule le_imp_inverse_le) show "0 < norm a - r" using r2 by simp next - have "norm a - norm (X i) \ norm (a - X i)" + have "norm a - norm (f x) \ norm (a - f x)" by (rule norm_triangle_ineq2) - also have "\ = norm (X i - a)" + also have "\ = norm (f x - a)" by (rule norm_minus_commute) also have "\ < r" using 1 . - finally show "norm a - r \ norm (X i)" by simp + finally show "norm a - r \ norm (f x)" by simp qed - finally show "norm (inverse (X i)) \ inverse (norm a - r)" . + finally show "norm (inverse (f x)) \ inverse (norm a - r)" . qed thus ?thesis by (rule BfunI) qed lemma tendsto_inverse_lemma: fixes a :: "'a::real_normed_div_algebra" - shows "\tendsto X a net; a \ 0; eventually (\i. X i \ 0) net\ - \ tendsto (\i. inverse (X i)) (inverse a) net" + shows "\(f ---> a) net; a \ 0; eventually (\x. f x \ 0) net\ + \ ((\x. inverse (f x)) ---> inverse a) net" apply (subst tendsto_Zfun_iff) apply (rule Zfun_ssubst) apply (erule eventually_elim1) @@ -530,23 +532,23 @@ lemma tendsto_inverse: fixes a :: "'a::real_normed_div_algebra" - assumes X: "tendsto X a net" + assumes f: "(f ---> a) net" assumes a: "a \ 0" - shows "tendsto (\i. inverse (X i)) (inverse a) net" + shows "((\x. inverse (f x)) ---> inverse a) net" proof - from a have "0 < norm a" by simp - with X have "eventually (\i. dist (X i) a < norm a) net" + with f have "eventually (\x. dist (f x) a < norm a) net" by (rule tendstoD) - then have "eventually (\i. X i \ 0) net" + then have "eventually (\x. f x \ 0) net" unfolding dist_norm by (auto elim!: eventually_elim1) - with X a show ?thesis + with f a show ?thesis by (rule tendsto_inverse_lemma) qed lemma tendsto_divide: fixes a b :: "'a::real_normed_field" - shows "\tendsto X a net; tendsto Y b net; b \ 0\ - \ tendsto (\n. X n / Y n) (a / b) net" + shows "\(f ---> a) net; (g ---> b) net; b \ 0\ + \ ((\x. f x / g x) ---> a / b) net" by (simp add: mult.tendsto tendsto_inverse divide_inverse) end diff -r 40f815edbde4 -r 93938cafc0e6 src/HOL/SEQ.thy --- a/src/HOL/SEQ.thy Fri Jun 05 21:55:08 2009 +0200 +++ b/src/HOL/SEQ.thy Fri Jun 05 15:59:20 2009 -0700 @@ -193,7 +193,7 @@ subsection {* Limits of Sequences *} -lemma LIMSEQ_conv_tendsto: "(X ----> L) \ tendsto X L sequentially" +lemma LIMSEQ_conv_tendsto: "(X ----> L) \ (X ---> L) sequentially" unfolding LIMSEQ_def tendsto_def eventually_sequentially .. lemma LIMSEQ_iff: