# HG changeset patch # User huffman # Date 1244011786 25200 # Node ID e23dd3aac0fb6682a057cf2f0eb07bc04b30d124 # Parent 1f72869f1a2e4a77624001fe81c39a3cc757f027 instance ^ :: complete_space diff -r 1f72869f1a2e -r e23dd3aac0fb src/HOL/Library/Euclidean_Space.thy --- a/src/HOL/Library/Euclidean_Space.thy Tue Jun 02 23:35:52 2009 -0700 +++ b/src/HOL/Library/Euclidean_Space.thy Tue Jun 02 23:49:46 2009 -0700 @@ -553,26 +553,8 @@ unfolding LIM_conv_tendsto by (rule tendsto_Cart_nth) lemma Cauchy_Cart_nth: - fixes X :: "nat \ 'a::metric_space ^ 'n::finite" - assumes "Cauchy (\n. X n)" - shows "Cauchy (\n. X n $ i)" -proof (rule metric_CauchyI) - fix e :: real assume "0 < e" - obtain M where "\m\M. \n\M. dist (X m) (X n) < e" - using metric_CauchyD [OF `Cauchy X` `0 < e`] by fast - moreover - { - fix m n - assume "M \ m" "M \ n" - have "dist (X m $ i) (X n $ i) \ dist (X m) (X n)" - by (rule dist_nth_le) - also assume "dist (X m) (X n) < e" - finally have "dist (X m $ i) (X n $ i) < e" . - } - ultimately - have "\m\M. \n\M. dist (X m $ i) (X n $ i) < e" by fast - thus "\M. \m\M. \n\M. dist (X m $ i) (X n $ i) < e" .. -qed + "Cauchy (\n. X n) \ Cauchy (\n. X n $ i)" +unfolding Cauchy_def by (fast intro: le_less_trans [OF dist_nth_le]) lemma LIMSEQ_vector: fixes X :: "nat \ 'a::metric_space ^ 'n::finite" @@ -641,6 +623,18 @@ then show "\M. \m\M. \n\M. dist (X m) (X n) < r" .. qed +instance "^" :: (complete_space, finite) complete_space +proof + fix X :: "nat \ 'a ^ 'b" assume "Cauchy X" + have "\i. (\n. X n $ i) ----> lim (\n. X n $ i)" + using Cauchy_Cart_nth [OF `Cauchy X`] + by (simp add: Cauchy_convergent_iff convergent_LIMSEQ_iff) + hence "X ----> Cart_lambda (\i. lim (\n. X n $ i))" + by (simp add: LIMSEQ_vector) + then show "convergent X" + by (rule convergentI) +qed + subsection {* Norms *} instantiation "^" :: (real_normed_vector, finite) real_normed_vector @@ -688,6 +682,8 @@ apply (rule_tac x="1" in exI, simp add: norm_nth_le) done +instance "^" :: (banach, finite) banach .. + subsection {* Inner products *} instantiation "^" :: (real_inner, finite) real_inner