diff -r 4d10e7f342b1 -r 6fb54701a11b src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space.thy --- a/src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space.thy Mon Aug 15 12:18:34 2011 -0700 +++ b/src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space.thy Mon Aug 15 14:09:39 2011 -0700 @@ -1093,34 +1093,21 @@ done qed +lemma continuous_at_component: "continuous (at a) (\x. x $ i)" +unfolding continuous_at by (intro tendsto_intros) + +lemma continuous_on_component: "continuous_on s (\x. x $ i)" +unfolding continuous_on_def by (intro ballI tendsto_intros) + lemma closed_positive_orthant: "closed {x::real^'n. \i. 0 \x$i}" -proof- - let ?U = "UNIV :: 'n set" - let ?O = "{x::real^'n. \i. x$i\0}" - {fix x:: "real^'n" and i::'n assume H: "\e>0. \x'\?O. x' \ x \ dist x' x < e" - and xi: "x$i < 0" - from xi have th0: "-x$i > 0" by arith - from H[rule_format, OF th0] obtain x' where x': "x' \?O" "x' \ x" "dist x' x < -x $ i" by blast - have th:" \b a (x::real). abs x <= b \ b <= a ==> ~(a + x < 0)" by arith - have th': "\x (y::real). x < 0 \ 0 <= y ==> abs x <= abs (y - x)" by arith - have th1: "\x$i\ \ \(x' - x)$i\" using x'(1) xi - apply (simp only: vector_component) - by (rule th') auto - have th2: "\dist x x'\ \ \(x' - x)$i\" using component_le_norm_cart[of "x'-x" i] - apply (simp add: dist_norm) by norm - from th[OF th1 th2] x'(3) have False by (simp add: dist_commute) } - then show ?thesis unfolding closed_limpt islimpt_approachable - unfolding not_le[symmetric] by blast -qed + unfolding Collect_all_eq + by (intro closed_INT ballI closed_Collect_le continuous_const + continuous_at_component) + lemma Lim_component_cart: fixes f :: "'a \ 'b::metric_space ^ 'n" shows "(f ---> l) net \ ((\a. f a $i) ---> l$i) net" - unfolding tendsto_iff - apply (clarify) - apply (drule spec, drule (1) mp) - apply (erule eventually_elim1) - apply (erule le_less_trans [OF dist_vec_nth_le]) - done + by (intro tendsto_intros) lemma bounded_component_cart: "bounded s \ bounded ((\x. x $ i) ` s)" unfolding bounded_def @@ -1193,12 +1180,6 @@ with r show "\l r. subseq r \ ((f \ r) ---> l) sequentially" by auto qed -lemma continuous_at_component: "continuous (at a) (\x. x $ i)" -unfolding continuous_at by (intro tendsto_intros) - -lemma continuous_on_component: "continuous_on s (\x. x $ i)" -unfolding continuous_on_def by (intro ballI tendsto_intros) - lemma interval_cart: fixes a :: "'a::ord^'n" shows "{a <..< b} = {x::'a^'n. \i. a$i < x$i \ x$i < b$i}" and "{a .. b} = {x::'a^'n. \i. a$i \ x$i \ x$i \ b$i}" @@ -1307,27 +1288,15 @@ lemma closed_interval_left_cart: fixes b::"real^'n" shows "closed {x::real^'n. \i. x$i \ b$i}" -proof- - { fix i - fix x::"real^'n" assume x:"\e>0. \x'\{x. \i. x $ i \ b $ i}. x' \ x \ dist x' x < e" - { assume "x$i > b$i" - then obtain y where "y $ i \ b $ i" "y \ x" "dist y x < x$i - b$i" using x[THEN spec[where x="x$i - b$i"]] by auto - hence False using component_le_norm_cart[of "y - x" i] unfolding dist_norm and vector_minus_component by auto } - hence "x$i \ b$i" by(rule ccontr)auto } - thus ?thesis unfolding closed_limpt unfolding islimpt_approachable by blast -qed + unfolding Collect_all_eq + by (intro closed_INT ballI closed_Collect_le continuous_const + continuous_at_component) lemma closed_interval_right_cart: fixes a::"real^'n" shows "closed {x::real^'n. \i. a$i \ x$i}" -proof- - { fix i - fix x::"real^'n" assume x:"\e>0. \x'\{x. \i. a $ i \ x $ i}. x' \ x \ dist x' x < e" - { assume "a$i > x$i" - then obtain y where "a $ i \ y $ i" "y \ x" "dist y x < a$i - x$i" using x[THEN spec[where x="a$i - x$i"]] by auto - hence False using component_le_norm_cart[of "y - x" i] unfolding dist_norm and vector_minus_component by auto } - hence "a$i \ x$i" by(rule ccontr)auto } - thus ?thesis unfolding closed_limpt unfolding islimpt_approachable by blast -qed + unfolding Collect_all_eq + by (intro closed_INT ballI closed_Collect_le continuous_const + continuous_at_component) lemma is_interval_cart:"is_interval (s::(real^'n) set) \ (\a\s. \b\s. \x. (\i. ((a$i \ x$i \ x$i \ b$i) \ (b$i \ x$i \ x$i \ a$i))) \ x \ s)" @@ -1335,19 +1304,19 @@ lemma closed_halfspace_component_le_cart: shows "closed {x::real^'n. x$i \ a}" - using closed_halfspace_le[of "(cart_basis i)::real^'n" a] unfolding inner_basis[OF assms] by auto + by (intro closed_Collect_le continuous_at_component continuous_const) lemma closed_halfspace_component_ge_cart: shows "closed {x::real^'n. x$i \ a}" - using closed_halfspace_ge[of a "(cart_basis i)::real^'n"] unfolding inner_basis[OF assms] by auto + by (intro closed_Collect_le continuous_at_component continuous_const) lemma open_halfspace_component_lt_cart: shows "open {x::real^'n. x$i < a}" - using open_halfspace_lt[of "(cart_basis i)::real^'n" a] unfolding inner_basis[OF assms] by auto + by (intro open_Collect_less continuous_at_component continuous_const) lemma open_halfspace_component_gt_cart: shows "open {x::real^'n. x$i > a}" - using open_halfspace_gt[of a "(cart_basis i)::real^'n"] unfolding inner_basis[OF assms] by auto + by (intro open_Collect_less continuous_at_component continuous_const) lemma Lim_component_le_cart: fixes f :: "'a \ real^'n" assumes "(f ---> l) net" "\ (trivial_limit net)" "eventually (\x. f(x)$i \ b) net" @@ -1382,23 +1351,14 @@ unfolding subspace_def by auto lemma closed_substandard_cart: - "closed {x::real^'n. \i. P i --> x$i = 0}" (is "closed ?A") + "closed {x::'a::real_normed_vector ^ 'n. \i. P i \ x$i = 0}" proof- - let ?D = "{i. P i}" - let ?Bs = "{{x::real^'n. inner (cart_basis i) x = 0}| i. i \ ?D}" - { fix x - { assume "x\?A" - hence x:"\i\?D. x $ i = 0" by auto - hence "x\ \ ?Bs" by(auto simp add: inner_basis x) } - moreover - { assume x:"x\\?Bs" - { fix i assume i:"i \ ?D" - then obtain B where BB:"B \ ?Bs" and B:"B = {x::real^'n. inner (cart_basis i) x = 0}" by auto - hence "x $ i = 0" unfolding B using x unfolding inner_basis by auto } - hence "x\?A" by auto } - ultimately have "x\?A \ x\ \?Bs" .. } - hence "?A = \ ?Bs" by auto - thus ?thesis by(auto simp add: closed_Inter closed_hyperplane) + { fix i::'n + have "closed {x::'a ^ 'n. P i \ x$i = 0}" + by (cases "P i", simp_all, intro closed_Collect_eq + continuous_at_component continuous_const) } + thus ?thesis + unfolding Collect_all_eq by (simp add: closed_INT) qed lemma dim_substandard_cart: