# HG changeset patch # User huffman # Date 1313611811 25200 # Node ID 336dd390e4a4f1b5d1b518e55b36c5bf49c4f9db # Parent c073a0bd84584cdc82064d5e7227d6386f2c865e Lim.thy: generalize and simplify proofs of LIM/LIMSEQ theorems diff -r c073a0bd8458 -r 336dd390e4a4 src/HOL/Lim.thy --- a/src/HOL/Lim.thy Wed Aug 17 11:39:09 2011 -0700 +++ b/src/HOL/Lim.thy Wed Aug 17 13:10:11 2011 -0700 @@ -543,80 +543,49 @@ subsection {* Relation of LIM and LIMSEQ *} lemma LIMSEQ_SEQ_conv1: - assumes X: "X -- a --> L" - shows "\S. (\n. S n \ a) \ S ----> a \ (\n. X (S n)) ----> L" - using tendsto_compose_eventually [OF X, where F=sequentially] by simp + fixes f :: "'a::topological_space \ 'b::topological_space" + assumes f: "f -- a --> l" + shows "\S. (\n. S n \ a) \ S ----> a \ (\n. f (S n)) ----> l" + using tendsto_compose_eventually [OF f, where F=sequentially] by simp + +lemma LIMSEQ_SEQ_conv2_lemma: + fixes a :: "'a::metric_space" + assumes "\S. (\n. S n \ a) \ S ----> a \ eventually (\x. P (S x)) sequentially" + shows "eventually P (at a)" +proof (rule ccontr) + let ?I = "\n. inverse (real (Suc n))" + let ?F = "\n::nat. SOME x. x \ a \ dist x a < ?I n \ \ P x" + assume "\ eventually P (at a)" + hence P: "\d>0. \x. x \ a \ dist x a < d \ \ P x" + unfolding eventually_at by simp + hence "\n. \x. x \ a \ dist x a < ?I n \ \ P x" by simp + hence F: "\n. ?F n \ a \ dist (?F n) a < ?I n \ \ P (?F n)" + by (rule someI_ex) + hence F1: "\n. ?F n \ a" + and F2: "\n. dist (?F n) a < ?I n" + and F3: "\n. \ P (?F n)" + by fast+ + have "?F ----> a" + using LIMSEQ_inverse_real_of_nat + by (rule metric_tendsto_imp_tendsto, + simp add: dist_norm F2 [THEN less_imp_le]) + moreover have "\n. ?F n \ a" + by (rule allI) (rule F1) + ultimately have "eventually (\n. P (?F n)) sequentially" + using assms by simp + thus "False" by (simp add: F3) +qed lemma LIMSEQ_SEQ_conv2: - fixes a :: real and L :: "'a::metric_space" - assumes "\S. (\n. S n \ a) \ S ----> a \ (\n. X (S n)) ----> L" - shows "X -- a --> L" -proof (rule ccontr) - assume "\ (X -- a --> L)" - hence "\ (\r > 0. \s > 0. \x. x \ a & norm (x - a) < s --> dist (X x) L < r)" - unfolding LIM_def dist_norm . - hence "\r > 0. \s > 0. \x. \(x \ a \ \x - a\ < s --> dist (X x) L < r)" by simp - hence "\r > 0. \s > 0. \x. (x \ a \ \x - a\ < s \ dist (X x) L \ r)" by (simp add: not_less) - then obtain r where rdef: "r > 0 \ (\s > 0. \x. (x \ a \ \x - a\ < s \ dist (X x) L \ r))" by auto - - let ?F = "\n::nat. SOME x. x\a \ \x - a\ < inverse (real (Suc n)) \ dist (X x) L \ r" - have "\n. \x. x\a \ \x - a\ < inverse (real (Suc n)) \ dist (X x) L \ r" - using rdef by simp - hence F: "\n. ?F n \ a \ \?F n - a\ < inverse (real (Suc n)) \ dist (X (?F n)) L \ r" - by (rule someI_ex) - hence F1: "\n. ?F n \ a" - and F2: "\n. \?F n - a\ < inverse (real (Suc n))" - and F3: "\n. dist (X (?F n)) L \ r" - by fast+ - - have "?F ----> a" - proof (rule LIMSEQ_I, unfold real_norm_def) - fix e::real - assume "0 < e" - (* choose no such that inverse (real (Suc n)) < e *) - then have "\no. inverse (real (Suc no)) < e" by (rule reals_Archimedean) - then obtain m where nodef: "inverse (real (Suc m)) < e" by auto - show "\no. \n. no \ n --> \?F n - a\ < e" - proof (intro exI allI impI) - fix n - assume mlen: "m \ n" - have "\?F n - a\ < inverse (real (Suc n))" - by (rule F2) - also have "inverse (real (Suc n)) \ inverse (real (Suc m))" - using mlen by auto - also from nodef have - "inverse (real (Suc m)) < e" . - finally show "\?F n - a\ < e" . - qed - qed - - moreover have "\n. ?F n \ a" - by (rule allI) (rule F1) - - moreover note `\S. (\n. S n \ a) \ S ----> a \ (\n. X (S n)) ----> L` - ultimately have "(\n. X (?F n)) ----> L" by simp - - moreover have "\ ((\n. X (?F n)) ----> L)" - proof - - { - fix no::nat - obtain n where "n = no + 1" by simp - then have nolen: "no \ n" by simp - (* We prove this by showing that for any m there is an n\m such that |X (?F n) - L| \ r *) - have "dist (X (?F n)) L \ r" - by (rule F3) - with nolen have "\n. no \ n \ dist (X (?F n)) L \ r" by fast - } - then have "(\no. \n. no \ n \ dist (X (?F n)) L \ r)" by simp - with rdef have "\e>0. (\no. \n. no \ n \ dist (X (?F n)) L \ e)" by auto - thus ?thesis by (unfold LIMSEQ_def, auto simp add: not_less) - qed - ultimately show False by simp -qed + fixes f :: "'a::metric_space \ 'b::topological_space" + assumes "\S. (\n. S n \ a) \ S ----> a \ (\n. f (S n)) ----> l" + shows "f -- a --> l" + using assms unfolding tendsto_def [where l=l] + by (simp add: LIMSEQ_SEQ_conv2_lemma) lemma LIMSEQ_SEQ_conv: - "(\S. (\n. S n \ a) \ S ----> (a::real) \ (\n. X (S n)) ----> L) = - (X -- a --> (L::'a::metric_space))" + "(\S. (\n. S n \ a) \ S ----> (a::'a::metric_space) \ (\n. X (S n)) ----> L) = + (X -- a --> (L::'b::topological_space))" using LIMSEQ_SEQ_conv2 LIMSEQ_SEQ_conv1 .. end