# HG changeset patch # User hoelzl # Date 1355496361 -3600 # Node ID 899c9c4e4a4c0a212021b69897bf054265167d91 # Parent 46be26e02456342e91f66baba1fd325515b7d69a Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors diff -r 46be26e02456 -r 899c9c4e4a4c NEWS --- a/NEWS Fri Dec 14 14:46:01 2012 +0100 +++ b/NEWS Fri Dec 14 15:46:01 2012 +0100 @@ -227,6 +227,36 @@ * HOL/Cardinals: Theories of ordinals and cardinals (supersedes the AFP entry "Ordinals_and_Cardinals"). +* HOL/Multivariate_Analysis: + Replaced "basis :: 'a::euclidean_space => nat => real" and + "\\ :: (nat => real) => 'a::euclidean_space" on euclidean spaces by + using the inner product "_ \ _" with vectors from the Basis set. + "\\ i. f i" is replaced by "SUM i : Basis. f i *r i". + + With this change the following constants are also chanegd or removed: + + DIM('a) :: nat ~> card (Basis :: 'a set) (is an abbreviation) + a $$ i ~> inner a i (where i : Basis) + cart_base i removed + \, \' removed + + Theorems about these constants where removed. + + Renamed lemmas: + + component_le_norm ~> Basis_le_norm + euclidean_eq ~> euclidean_eq_iff + differential_zero_maxmin_component ~> differential_zero_maxmin_cart + euclidean_simps ~> inner_simps + independent_basis ~> independent_Basis + span_basis ~> span_Basis + in_span_basis ~> in_span_Basis + norm_bound_component_le ~> norm_boound_Basis_le + norm_bound_component_lt ~> norm_boound_Basis_lt + component_le_infnorm ~> Basis_le_infnorm + + INCOMPATIBILITY. + * HOL/Probability: - Add simproc "measurable" to automatically prove measurability diff -r 46be26e02456 -r 899c9c4e4a4c src/HOL/Multivariate_Analysis/Brouwer_Fixpoint.thy --- a/src/HOL/Multivariate_Analysis/Brouwer_Fixpoint.thy Fri Dec 14 14:46:01 2012 +0100 +++ b/src/HOL/Multivariate_Analysis/Brouwer_Fixpoint.thy Fri Dec 14 15:46:01 2012 +0100 @@ -22,6 +22,18 @@ imports Convex_Euclidean_Space begin +(** move this **) +lemma divide_nonneg_nonneg:assumes "a \ 0" "b \ 0" shows "0 \ a / (b::real)" + apply(cases "b=0") defer apply(rule divide_nonneg_pos) using assms by auto + +lemma continuous_setsum: + fixes f :: "'i \ 'a::t2_space \ 'b::real_normed_vector" + assumes f: "\i. i \ I \ continuous F (f i)" shows "continuous F (\x. \i\I. f i x)" +proof cases + assume "finite I" from this f show ?thesis + by (induct I) (auto intro!: continuous_intros) +qed (auto intro!: continuous_intros) + lemma brouwer_compactness_lemma: assumes "compact s" "continuous_on s f" "\ (\x\s. (f x = (0::_::euclidean_space)))" obtains d where "0 < d" "\x\s. d \ norm(f x)" @@ -39,39 +51,39 @@ qed lemma kuhn_labelling_lemma: - fixes type::"'a::euclidean_space" - assumes "(\x::'a. P x \ P (f x))" - and "\x. P x \ (\i 0 \ x$$i \ x$$i \ 1)" - shows "\l. (\x.\ i (1::nat)) \ - (\x.\ i Q i \ (x$$i = 0) \ (l x i = 0)) \ - (\x.\ i Q i \ (x$$i = 1) \ (l x i = 1)) \ - (\x.\ i Q i \ (l x i = 0) \ x$$i \ f(x)$$i) \ - (\x.\ i Q i \ (l x i = 1) \ f(x)$$i \ x$$i)" + fixes P Q :: "'a::euclidean_space \ bool" + assumes "(\x. P x \ P (f x))" + and "\x. P x \ (\i\Basis. Q i \ 0 \ x\i \ x\i \ 1)" + shows "\l. (\x.\i\Basis. l x i \ (1::nat)) \ + (\x.\i\Basis. P x \ Q i \ (x\i = 0) \ (l x i = 0)) \ + (\x.\i\Basis. P x \ Q i \ (x\i = 1) \ (l x i = 1)) \ + (\x.\i\Basis. P x \ Q i \ (l x i = 0) \ x\i \ f(x)\i) \ + (\x.\i\Basis. P x \ Q i \ (l x i = 1) \ f(x)\i \ x\i)" proof - have and_forall_thm:"\P Q. (\x. P x) \ (\x. Q x) \ (\x. P x \ Q x)" by auto have *: "\x y::real. 0 \ x \ x \ 1 \ 0 \ y \ y \ 1 \ (x \ 1 \ x \ y \ x \ 0 \ y \ x)" by auto show ?thesis - unfolding and_forall_thm + unfolding and_forall_thm Ball_def apply(subst choice_iff[THEN sym])+ apply rule apply rule proof - - case goal1 - let ?R = "\y. (P x \ Q xa \ x $$ xa = 0 \ y = (0::nat)) \ - (P x \ Q xa \ x $$ xa = 1 \ y = 1) \ - (P x \ Q xa \ y = 0 \ x $$ xa \ f x $$ xa) \ - (P x \ Q xa \ y = 1 \ f x $$ xa \ x $$ xa)" + case (goal1 x) + let ?R = "\y. (P x \ Q xa \ x \ xa = 0 \ y = (0::nat)) \ + (P x \ Q xa \ x \ xa = 1 \ y = 1) \ + (P x \ Q xa \ y = 0 \ x \ xa \ f x \ xa) \ + (P x \ Q xa \ y = 1 \ f x \ xa \ x \ xa)" { - assume "P x" "Q xa" "xa f x $$ xa \ f x $$ xa \ 1" + assume "P x" "Q xa" "xa\Basis" + then have "0 \ f x \ xa \ f x \ xa \ 1" using assms(2)[rule_format,of "f x" xa] apply (drule_tac assms(1)[rule_format]) apply auto done } - then have "xa ?R 0 \ ?R 1" by auto + then have "xa\Basis \ ?R 0 \ ?R 1" by auto then show ?case by auto qed qed @@ -1363,50 +1375,56 @@ apply(drule_tac assms(1)[rule_format]) by auto } hence "?R 0 \ ?R 1" by auto thus ?case by auto qed qed -lemma brouwer_cube: fixes f::"'a::ordered_euclidean_space \ 'a::ordered_euclidean_space" - assumes "continuous_on {0..\\ i. 1} f" "f ` {0..\\ i. 1} \ {0..\\ i. 1}" - shows "\x\{0..\\ i. 1}. f x = x" apply(rule ccontr) proof- - def n \ "DIM('a)" have n:"1 \ n" "0 < n" "n \ 0" unfolding n_def by(auto simp add:Suc_le_eq) - assume "\ (\x\{0..\\ i. 1}. f x = x)" hence *:"\ (\x\{0..\\ i. 1}. f x - x = 0)" by auto +lemma brouwer_cube: + fixes f::"'a::ordered_euclidean_space \ 'a::ordered_euclidean_space" + assumes "continuous_on {0..(\Basis)} f" "f ` {0..(\Basis)} \ {0..(\Basis)}" + shows "\x\{0..(\Basis)}. f x = x" + proof (rule ccontr) + def n \ "DIM('a)" have n:"1 \ n" "0 < n" "n \ 0" unfolding n_def by(auto simp add: Suc_le_eq DIM_positive) + assume "\ (\x\{0..\Basis}. f x = x)" hence *:"\ (\x\{0..\Basis}. f x - x = 0)" by auto guess d apply(rule brouwer_compactness_lemma[OF compact_interval _ *]) apply(rule continuous_on_intros assms)+ . note d=this[rule_format] - have *:"\x. x \ {0..\\ i. 1} \ f x \ {0..\\ i. 1}" "\x. x \ {0..(\\ i. 1)::'a} \ - (\i 0 \ x $$ i \ x $$ i \ 1)" + have *:"\x. x \ {0..\Basis} \ f x \ {0..\Basis}" "\x. x \ {0..(\Basis)::'a} \ + (\i\Basis. True \ 0 \ x \ i \ x \ i \ 1)" using assms(2)[unfolded image_subset_iff Ball_def] unfolding mem_interval by auto guess label using kuhn_labelling_lemma[OF *] apply-apply(erule exE,(erule conjE)+) . note label = this[rule_format] - have lem1:"\x\{0..\\ i. 1}.\y\{0..\\ i. 1}.\i label y i - \ abs(f x $$ i - x $$ i) \ norm(f y - f x) + norm(y - x)" proof safe - fix x y::'a assume xy:"x\{0..\\ i. 1}" "y\{0..\\ i. 1}" fix i assume i:"label x i \ label y i" "ix\{0..\Basis}.\y\{0..\Basis}.\i\Basis. label x i \ label y i + \ abs(f x \ i - x \ i) \ norm(f y - f x) + norm(y - x)" proof safe + fix x y::'a assume xy:"x\{0..\Basis}" "y\{0..\Basis}" + fix i assume i:"label x i \ label y i" "i\Basis" have *:"\x y fx fy::real. (x \ fx \ fy \ y \ fx \ x \ y \ fy) \ abs(fx - x) \ abs(fy - fx) + abs(y - x)" by auto - have "\(f x - x) $$ i\ \ abs((f y - f x)$$i) + abs((y - x)$$i)" unfolding euclidean_simps + have "\(f x - x) \ i\ \ abs((f y - f x)\i) + abs((y - x)\i)" + unfolding inner_simps apply(rule *) apply(cases "label x i = 0") apply(rule disjI1,rule) prefer 3 proof(rule disjI2,rule) assume lx:"label x i = 0" hence ly:"label y i = 1" using i label(1)[of i y] by auto - show "x $$ i \ f x $$ i" apply(rule label(4)[rule_format]) using xy lx i(2) by auto - show "f y $$ i \ y $$ i" apply(rule label(5)[rule_format]) using xy ly i(2) by auto next + show "x \ i \ f x \ i" apply(rule label(4)[rule_format]) using xy lx i(2) by auto + show "f y \ i \ y \ i" apply(rule label(5)[rule_format]) using xy ly i(2) by auto next assume "label x i \ 0" hence l:"label x i = 1" "label y i = 0" using i label(1)[of i x] label(1)[of i y] by auto - show "f x $$ i \ x $$ i" apply(rule label(5)[rule_format]) using xy l i(2) by auto - show "y $$ i \ f y $$ i" apply(rule label(4)[rule_format]) using xy l i(2) by auto qed - also have "\ \ norm (f y - f x) + norm (y - x)" apply(rule add_mono) by(rule component_le_norm)+ - finally show "\f x $$ i - x $$ i\ \ norm (f y - f x) + norm (y - x)" unfolding euclidean_simps . qed - have "\e>0. \x\{0..\\ i. 1}. \y\{0..\\ i. 1}. \z\{0..\\ i. 1}. \i norm(y - z) < e \ label x i \ label y i \ abs((f(z) - z)$$i) < d / (real n)" proof- - have d':"d / real n / 8 > 0" apply(rule divide_pos_pos)+ using d(1) unfolding n_def by auto - have *:"uniformly_continuous_on {0..\\ i. 1} f" by(rule compact_uniformly_continuous[OF assms(1) compact_interval]) + show "f x \ i \ x \ i" apply(rule label(5)[rule_format]) using xy l i(2) by auto + show "y \ i \ f y \ i" apply(rule label(4)[rule_format]) using xy l i(2) by auto qed + also have "\ \ norm (f y - f x) + norm (y - x)" apply(rule add_mono) by(rule Basis_le_norm[OF i(2)])+ + finally show "\f x \ i - x \ i\ \ norm (f y - f x) + norm (y - x)" unfolding inner_simps . + qed + have "\e>0. \x\{0..\Basis}. \y\{0..\Basis}. \z\{0..\Basis}. \i\Basis. + norm(x - z) < e \ norm(y - z) < e \ label x i \ label y i \ abs((f(z) - z)\i) < d / (real n)" proof- + have d':"d / real n / 8 > 0" apply(rule divide_pos_pos)+ using d(1) unfolding n_def by (auto simp: DIM_positive) + have *:"uniformly_continuous_on {0..\Basis} f" by(rule compact_uniformly_continuous[OF assms(1) compact_interval]) guess e using *[unfolded uniformly_continuous_on_def,rule_format,OF d'] apply-apply(erule exE,(erule conjE)+) . note e=this[rule_format,unfolded dist_norm] show ?thesis apply(rule_tac x="min (e/2) (d/real n/8)" in exI) - proof safe show "0 < min (e / 2) (d / real n / 8)" using d' e by auto - fix x y z i assume as:"x \ {0..\\ i. 1}" "y \ {0..\\ i. 1}" "z \ {0..\\ i. 1}" + proof safe + show "0 < min (e / 2) (d / real n / 8)" using d' e by auto + fix x y z i assume as:"x \ {0..\Basis}" "y \ {0..\Basis}" "z \ {0..\Basis}" "norm (x - z) < min (e / 2) (d / real n / 8)" - "norm (y - z) < min (e / 2) (d / real n / 8)" "label x i \ label y i" and i:"i label y i" and i:"i\Basis" have *:"\z fz x fx n1 n2 n3 n4 d4 d::real. abs(fx - x) \ n1 + n2 \ abs(fx - fz) \ n3 \ abs(x - z) \ n4 \ n1 < d4 \ n2 < 2 * d4 \ n3 < d4 \ n4 < d4 \ (8 * d4 = d) \ abs(fz - z) < d" by auto - show "\(f z - z) $$ i\ < d / real n" unfolding euclidean_simps proof(rule *) - show "\f x $$ i - x $$ i\ \ norm (f y -f x) + norm (y - x)" apply(rule lem1[rule_format]) using as i by auto - show "\f x $$ i - f z $$ i\ \ norm (f x - f z)" "\x $$ i - z $$ i\ \ norm (x - z)" - unfolding euclidean_component_diff[THEN sym] by(rule component_le_norm)+ + show "\(f z - z) \ i\ < d / real n" unfolding inner_simps proof(rule *) + show "\f x \ i - x \ i\ \ norm (f y -f x) + norm (y - x)" apply(rule lem1[rule_format]) using as i by auto + show "\f x \ i - f z \ i\ \ norm (f x - f z)" "\x \ i - z \ i\ \ norm (x - z)" + unfolding inner_diff_left[THEN sym] by(rule Basis_le_norm[OF i])+ have tria:"norm (y - x) \ norm (y - z) + norm (x - z)" using dist_triangle[of y x z,unfolded dist_norm] unfolding norm_minus_commute by auto also have "\ < e / 2 + e / 2" apply(rule add_strict_mono) using as(4,5) by auto @@ -1418,95 +1436,99 @@ guess p using real_arch_simple[of "1 + real n / e"] .. note p=this have "1 + real n / e > 0" apply(rule add_pos_pos) defer apply(rule divide_pos_pos) using e(1) n by auto hence "p > 0" using p by auto - def b \ "\i. i - 1::nat" have b:"bij_betw b {1..n} {.. 'a" where b: "bij_betw b {1..n} Basis" + by atomize_elim (auto simp: n_def intro!: finite_same_card_bij) def b' \ "inv_into {1..n} b" - have b':"bij_betw b' {..i. i b (b' i) = i" unfolding b'_def apply(rule f_inv_into_f) using b - unfolding bij_betw_def by auto - have b'b[simp]:"\i. i\{1..n} \ b' (b i) = i" unfolding b'_def apply(rule inv_into_f_eq) - using b unfolding n_def bij_betw_def by auto + then have b': "bij_betw b' Basis {1..n}" + using bij_betw_inv_into[OF b] by auto + then have b'_Basis: "\i. i \ Basis \ b' i \ {Suc 0 .. n}" + unfolding bij_betw_def by (auto simp: set_eq_iff) + have bb'[simp]:"\i. i \ Basis \ b (b' i) = i" + unfolding b'_def using b by (auto simp: f_inv_into_f bij_betw_def) + have b'b[simp]:"\i. i \ {1..n} \ b' (b i) = i" + unfolding b'_def using b by (auto simp: inv_into_f_eq bij_betw_def) have *:"\x::nat. x=0 \ x=1 \ x\1" by auto - have b'':"\j. j\{1..n} \ b j j. j\{Suc 0..n} \ b j \Basis" using b unfolding bij_betw_def by auto have q1:"0 < p" "0 < n" "\x. (\i\{1..n}. x i \ p) \ - (\i\{1..n}. (label (\\ i. real (x (b' i)) / real p) \ b) i = 0 \ (label (\\ i. real (x (b' i)) / real p) \ b) i = 1)" + (\i\{1..n}. (label (\i\Basis. (real (x (b' i)) / real p) *\<^sub>R i) \ b) i = 0 \ + (label (\i\Basis. (real (x (b' i)) / real p) *\<^sub>R i) \ b) i = 1)" unfolding * using `p>0` `n>0` using label(1)[OF b''] by auto - have q2:"\x. (\i\{1..n}. x i \ p) \ (\i\{1..n}. x i = 0 \ (label (\\ i. real (x (b' i)) / real p) \ b) i = 0)" - "\x. (\i\{1..n}. x i \ p) \ (\i\{1..n}. x i = p \ (label (\\ i. real (x (b' i)) / real p) \ b) i = 1)" - apply(rule,rule,rule,rule) defer proof(rule,rule,rule,rule) fix x i - assume as:"\i\{1..n}. x i \ p" "i \ {1..n}" - { assume "x i = p \ x i = 0" - have "(\\ i. real (x (b' i)) / real p) \ {0::'a..\\ i. 1}" unfolding mem_interval - apply safe unfolding euclidean_lambda_beta euclidean_component_zero - proof (simp_all only: if_P) fix j assume j':"j {1..n}" using b' unfolding n_def bij_betw_def by auto - show "0 \ real (x (b' j)) / real p" - apply(rule divide_nonneg_pos) using `p>0` using as(1)[rule_format,OF j] by auto - show "real (x (b' j)) / real p \ 1" unfolding divide_le_eq_1 - using as(1)[rule_format,OF j] by auto qed } note cube=this - { assume "x i = p" thus "(label (\\ i. real (x (b' i)) / real p) \ b) i = 1" unfolding o_def - apply- apply(rule label(3)) apply(rule b'') using cube using as `p>0` - proof safe assume i:"i\{1..n}" - show "((\\ ia. real (x (b' ia)) / real (x i))::'a) $$ b i = 1" - unfolding euclidean_lambda_beta apply(subst if_P) apply(rule b''[OF i]) unfolding b'b[OF i] - unfolding `x i = p` using q1(1) by auto - qed auto } - { assume "x i = 0" thus "(label (\\ i. real (x (b' i)) / real p) \ b) i = 0" unfolding o_def - apply-apply(rule label(2)[OF b'']) using cube using as `p>0` - proof safe assume i:"i\{1..n}" - show "((\\ ia. real (x (b' ia)) / real p)::'a) $$ b i = 0" - unfolding euclidean_lambda_beta apply (subst if_P) apply(rule b''[OF i]) unfolding b'b[OF i] - unfolding `x i = 0` using q1(1) by auto - qed auto } + have q2:"\x. (\i\{1..n}. x i \ p) \ (\i\{1..n}. x i = 0 \ + (label (\i\Basis. (real (x (b' i)) / real p) *\<^sub>R i) \ b) i = 0)" + "\x. (\i\{1..n}. x i \ p) \ (\i\{1..n}. x i = p \ + (label (\i\Basis. (real (x (b' i)) / real p) *\<^sub>R i) \ b) i = 1)" + apply(rule,rule,rule,rule) + defer + proof(rule,rule,rule,rule) + fix x i assume as:"\i\{1..n}. x i \ p" "i \ {1..n}" + { assume "x i = p \ x i = 0" + have "(\i\Basis. (real (x (b' i)) / real p) *\<^sub>R i) \ {0::'a..\Basis}" + unfolding mem_interval using as b'_Basis + by (auto simp add: inner_simps bij_betw_def zero_le_divide_iff divide_le_eq_1) } + note cube=this + { assume "x i = p" thus "(label (\i\Basis. (real (x (b' i)) / real p) *\<^sub>R i) \ b) i = 1" + unfolding o_def using cube as `p>0` + by (intro label(3)) (auto simp add: b'') } + { assume "x i = 0" thus "(label (\i\Basis. (real (x (b' i)) / real p) *\<^sub>R i) \ b) i = 0" + unfolding o_def using cube as `p>0` + by (intro label(2)) (auto simp add: b'') } qed guess q apply(rule kuhn_lemma[OF q1 q2]) . note q=this - def z \ "(\\ i. real (q (b' i)) / real p)::'a" - have "\i abs((f z - z)$$i)" proof(rule ccontr) - have "\i {0..i {0..p}" apply-apply(rule,erule_tac x=i in allE) by auto - hence "z\{0..\\ i.1}" unfolding z_def mem_interval apply safe unfolding euclidean_lambda_beta - unfolding euclidean_component_zero apply (simp_all only: if_P) - apply(rule divide_nonneg_pos) using `p>0` unfolding divide_le_eq_1 by auto - hence d_fz_z:"d \ norm (f z - z)" apply(drule_tac d) . - case goal1 hence as:"\if z $$ i - z $$ i\ < d / real n" using `n>0` by(auto simp add:not_le) - have "norm (f z - z) \ (\if z $$ i - z $$ i\)" unfolding euclidean_component_diff[THEN sym] by(rule norm_le_l1) - also have "\ < (\i = d" unfolding real_eq_of_nat n_def using n using DIM_positive[where 'a='a] by auto + def z \ "(\i\Basis. (real (q (b' i)) / real p) *\<^sub>R i)::'a" + have "\i\Basis. d / real n \ abs((f z - z)\i)" + proof(rule ccontr) + have "\i\Basis. q (b' i) \ {0..p}" + using q(1) b' by (auto intro: less_imp_le simp: bij_betw_def) + hence "z\{0..\Basis}" + unfolding z_def mem_interval using b'_Basis + by (auto simp add: inner_simps bij_betw_def zero_le_divide_iff divide_le_eq_1) + hence d_fz_z:"d \ norm (f z - z)" by (rule d) + case goal1 + hence as:"\i\Basis. \f z \ i - z \ i\ < d / real n" + using `n>0` by(auto simp add: not_le inner_simps) + have "norm (f z - z) \ (\i\Basis. \f z \ i - z \ i\)" + unfolding inner_diff_left[symmetric] by(rule norm_le_l1) + also have "\ < (\(i::'a)\Basis. d / real n)" apply(rule setsum_strict_mono) using as by auto + also have "\ = d" using DIM_positive[where 'a='a] by (auto simp: real_eq_of_nat n_def) finally show False using d_fz_z by auto qed then guess i .. note i=this have *:"b' i \ {1..n}" using i using b'[unfolded bij_betw_def] by auto guess r using q(2)[rule_format,OF *] .. then guess s apply-apply(erule exE,(erule conjE)+) . note rs=this[rule_format] - have b'_im:"\i. i b' i \ {1..n}" using b' unfolding bij_betw_def by auto - def r' \ "(\\ i. real (r (b' i)) / real p)::'a" - have "\i. i r (b' i) \ p" apply(rule order_trans) apply(rule rs(1)[OF b'_im,THEN conjunct2]) + have b'_im:"\i. i\Basis \ b' i \ {1..n}" using b' unfolding bij_betw_def by auto + def r' \ "(\i\Basis. (real (r (b' i)) / real p) *\<^sub>R i)::'a" + have "\i. i\Basis \ r (b' i) \ p" apply(rule order_trans) apply(rule rs(1)[OF b'_im,THEN conjunct2]) using q(1)[rule_format,OF b'_im] by(auto simp add: Suc_le_eq) - hence "r' \ {0..\\ i. 1}" unfolding r'_def mem_interval apply safe unfolding euclidean_lambda_beta euclidean_component_zero - apply (simp only: if_P) - apply(rule divide_nonneg_pos) using rs(1)[OF b'_im] q(1)[rule_format,OF b'_im] `p>0` by auto - def s' \ "(\\ i. real (s (b' i)) / real p)::'a" - have "\i. i s (b' i) \ p" apply(rule order_trans) apply(rule rs(2)[OF b'_im,THEN conjunct2]) + hence "r' \ {0..\Basis}" + unfolding r'_def mem_interval using b'_Basis + by (auto simp add: inner_simps bij_betw_def zero_le_divide_iff divide_le_eq_1) + def s' \ "(\i\Basis. (real (s (b' i)) / real p) *\<^sub>R i)::'a" + have "\i. i\Basis \ s (b' i) \ p" apply(rule order_trans) apply(rule rs(2)[OF b'_im,THEN conjunct2]) using q(1)[rule_format,OF b'_im] by(auto simp add: Suc_le_eq) - hence "s' \ {0..\\ i.1}" unfolding s'_def mem_interval apply safe unfolding euclidean_lambda_beta euclidean_component_zero - apply (simp_all only: if_P) apply(rule divide_nonneg_pos) using rs(1)[OF b'_im] q(1)[rule_format,OF b'_im] `p>0` by auto - have "z\{0..\\ i.1}" unfolding z_def mem_interval apply safe unfolding euclidean_lambda_beta euclidean_component_zero - apply (simp_all only: if_P) apply(rule divide_nonneg_pos) using q(1)[rule_format,OF b'_im] `p>0` by(auto intro:less_imp_le) + hence "s' \ {0..\Basis}" + unfolding s'_def mem_interval using b'_Basis + by (auto simp add: inner_simps bij_betw_def zero_le_divide_iff divide_le_eq_1) + have "z\{0..\Basis}" + unfolding z_def mem_interval using b'_Basis q(1)[rule_format,OF b'_im] `p>0` + by (auto simp add: inner_simps bij_betw_def zero_le_divide_iff divide_le_eq_1 less_imp_le) have *:"\x. 1 + real x = real (Suc x)" by auto - { have "(\ireal (r (b' i)) - real (q (b' i))\) \ (\ii\Basis. \real (r (b' i)) - real (q (b' i))\) \ (\(i::'a)\Basis. 1)" apply(rule setsum_mono) using rs(1)[OF b'_im] by(auto simp add:* field_simps) - also have "\ < e * real p" using p `e>0` `p>0` unfolding n_def real_of_nat_def - by(auto simp add:field_simps) - finally have "(\ireal (r (b' i)) - real (q (b' i))\) < e * real p" . } moreover - { have "(\ireal (s (b' i)) - real (q (b' i))\) \ (\i < e * real p" using p `e>0` `p>0` + by(auto simp add: field_simps n_def real_of_nat_def) + finally have "(\i\Basis. \real (r (b' i)) - real (q (b' i))\) < e * real p" . } moreover + { have "(\i\Basis. \real (s (b' i)) - real (q (b' i))\) \ (\(i::'a)\Basis. 1)" apply(rule setsum_mono) using rs(2)[OF b'_im] by(auto simp add:* field_simps) - also have "\ < e * real p" using p `e>0` `p>0` unfolding n_def real_of_nat_def - by(auto simp add:field_simps) - finally have "(\ireal (s (b' i)) - real (q (b' i))\) < e * real p" . } ultimately - have "norm (r' - z) < e" "norm (s' - z) < e" unfolding r'_def s'_def z_def apply- - apply(rule_tac[!] le_less_trans[OF norm_le_l1]) using `p>0` - by(auto simp add:field_simps setsum_divide_distrib[THEN sym]) - hence "\(f z - z) $$ i\ < d / real n" apply-apply(rule e(2)[OF `r'\{0..\\ i.1}` `s'\{0..\\ i.1}` `z\{0..\\ i.1}`]) - using rs(3) unfolding r'_def[symmetric] s'_def[symmetric] o_def bb' using i by auto - thus False using i by auto qed + also have "\ < e * real p" using p `e>0` `p>0` + by(auto simp add: field_simps n_def real_of_nat_def) + finally have "(\i\Basis. \real (s (b' i)) - real (q (b' i))\) < e * real p" . } ultimately + have "norm (r' - z) < e" "norm (s' - z) < e" unfolding r'_def s'_def z_def using `p>0` + by (rule_tac[!] le_less_trans[OF norm_le_l1]) + (auto simp add: field_simps setsum_divide_distrib[symmetric] inner_diff_left) + hence "\(f z - z) \ i\ < d / real n" + using rs(3) i unfolding r'_def[symmetric] s'_def[symmetric] o_def bb' + by (intro e(2)[OF `r'\{0..\Basis}` `s'\{0..\Basis}` `z\{0..\Basis}`]) auto + thus False using i by auto +qed subsection {* Retractions. *} @@ -1551,12 +1573,14 @@ subsection {*So the Brouwer theorem for any set with nonempty interior. *} -lemma brouwer_weak: fixes f::"'a::ordered_euclidean_space \ 'a::ordered_euclidean_space" +lemma brouwer_weak: + fixes f::"'a::ordered_euclidean_space \ 'a::ordered_euclidean_space" assumes "compact s" "convex s" "interior s \ {}" "continuous_on s f" "f ` s \ s" obtains x where "x \ s" "f x = x" proof- - have *:"interior {0::'a..\\ i.1} \ {}" unfolding interior_closed_interval interval_eq_empty by auto - have *:"{0::'a..\\ i.1} homeomorphic s" using homeomorphic_convex_compact[OF convex_interval(1) compact_interval * assms(2,1,3)] . - have "\f. continuous_on {0::'a..\\ i.1} f \ f ` {0::'a..\\ i.1} \ {0::'a..\\ i.1} \ (\x\{0::'a..\\ i.1}. f x = x)" + have *:"interior {0::'a..\Basis} \ {}" unfolding interior_closed_interval interval_eq_empty by auto + have *:"{0::'a..\Basis} homeomorphic s" using homeomorphic_convex_compact[OF convex_interval(1) compact_interval * assms(2,1,3)] . + have "\f. continuous_on {0::'a..\Basis} f \ f ` {0::'a..\Basis} \ {0::'a..\Basis} \ + (\x\{0::'a..\Basis}. f x = x)" using brouwer_cube by auto thus ?thesis unfolding homeomorphic_fixpoint_property[OF *] apply(erule_tac x=f in allE) apply(erule impE) defer apply(erule bexE) apply(rule_tac x=y in that) using assms by auto qed @@ -1609,49 +1633,45 @@ subsection {*Bijections between intervals. *} -definition "interval_bij = (\ (a::'a,b::'a) (u::'a,v::'a) (x::'a::ordered_euclidean_space). - (\\ i. u$$i + (x$$i - a$$i) / (b$$i - a$$i) * (v$$i - u$$i))::'a)" +definition interval_bij :: "'a \ 'a \ 'a \ 'a \ 'a \ 'a::ordered_euclidean_space" where + "interval_bij \ \(a, b) (u, v) x. (\i\Basis. (u\i + (x\i - a\i) / (b\i - a\i) * (v\i - u\i)) *\<^sub>R i)" lemma interval_bij_affine: - "interval_bij (a,b) (u,v) = (\x. (\\ i. (v$$i - u$$i) / (b$$i - a$$i) * x$$i) + - (\\ i. u$$i - (v$$i - u$$i) / (b$$i - a$$i) * a$$i))" - apply rule apply(subst euclidean_eq,safe) unfolding euclidean_simps interval_bij_def euclidean_lambda_beta - by(auto simp add: field_simps add_divide_distrib[THEN sym]) + "interval_bij (a,b) (u,v) = (\x. (\i\Basis. ((v\i - u\i) / (b\i - a\i) * (x\i)) *\<^sub>R i) + + (\i\Basis. (u\i - (v\i - u\i) / (b\i - a\i) * (a\i)) *\<^sub>R i))" + by (auto simp: setsum_addf[symmetric] scaleR_add_left[symmetric] interval_bij_def fun_eq_iff + field_simps inner_simps add_divide_distrib[symmetric] intro!: setsum_cong) lemma continuous_interval_bij: "continuous (at x) (interval_bij (a,b::'a::ordered_euclidean_space) (u,v::'a))" - unfolding interval_bij_affine apply(rule continuous_intros) - apply(rule linear_continuous_at) unfolding linear_conv_bounded_linear[THEN sym] - unfolding linear_def euclidean_eq[where 'a='a] apply safe unfolding euclidean_lambda_beta prefer 3 - apply(rule continuous_intros) by(auto simp add:field_simps add_divide_distrib[THEN sym]) + by (auto simp add: divide_inverse interval_bij_def intro!: continuous_setsum continuous_intros) lemma continuous_on_interval_bij: "continuous_on s (interval_bij (a,b) (u,v))" apply(rule continuous_at_imp_continuous_on) by(rule, rule continuous_interval_bij) -(** move this **) -lemma divide_nonneg_nonneg:assumes "a \ 0" "b \ 0" shows "0 \ a / (b::real)" - apply(cases "b=0") defer apply(rule divide_nonneg_pos) using assms by auto - -lemma in_interval_interval_bij: assumes "x \ {a..b}" "{u..v} \ {}" - shows "interval_bij (a,b) (u,v) x \ {u..v::'a::ordered_euclidean_space}" - unfolding interval_bij_def split_conv mem_interval apply safe unfolding euclidean_lambda_beta -proof (simp_all only: if_P) - fix i assume i:"i {}" using assms by auto - hence *:"a$$i \ b$$i" "u$$i \ v$$i" using assms(2) unfolding interval_eq_empty not_ex apply- - apply(erule_tac[!] x=i in allE)+ by auto - have x:"a$$i\x$$i" "x$$i\b$$i" using assms(1)[unfolded mem_interval] using i by auto - have "0 \ (x $$ i - a $$ i) / (b $$ i - a $$ i) * (v $$ i - u $$ i)" - apply(rule mult_nonneg_nonneg) apply(rule divide_nonneg_nonneg) - using * x by(auto simp add:field_simps) - thus "u $$ i \ u $$ i + (x $$ i - a $$ i) / (b $$ i - a $$ i) * (v $$ i - u $$ i)" using * by auto - have "((x $$ i - a $$ i) / (b $$ i - a $$ i)) * (v $$ i - u $$ i) \ 1 * (v $$ i - u $$ i)" +lemma in_interval_interval_bij: + fixes a b u v x :: "'a::ordered_euclidean_space" + assumes "x \ {a..b}" "{u..v} \ {}" shows "interval_bij (a,b) (u,v) x \ {u..v}" + apply (simp only: interval_bij_def split_conv mem_interval inner_setsum_left_Basis cong: ball_cong) +proof safe + fix i :: 'a assume i:"i\Basis" + have "{a..b} \ {}" using assms by auto + with i have *: "a\i \ b\i" "u\i \ v\i" + using assms(2) by (auto simp add: interval_eq_empty not_less) + have x: "a\i\x\i" "x\i\b\i" + using assms(1)[unfolded mem_interval] using i by auto + have "0 \ (x \ i - a \ i) / (b \ i - a \ i) * (v \ i - u \ i)" + using * x by (auto intro!: mult_nonneg_nonneg divide_nonneg_nonneg) + thus "u \ i \ u \ i + (x \ i - a \ i) / (b \ i - a \ i) * (v \ i - u \ i)" + using * by auto + have "((x \ i - a \ i) / (b \ i - a \ i)) * (v \ i - u \ i) \ 1 * (v \ i - u \ i)" apply(rule mult_right_mono) unfolding divide_le_eq_1 using * x by auto - thus "u $$ i + (x $$ i - a $$ i) / (b $$ i - a $$ i) * (v $$ i - u $$ i) \ v $$ i" using * by auto + thus "u \ i + (x \ i - a \ i) / (b \ i - a \ i) * (v \ i - u \ i) \ v \ i" using * by auto qed -lemma interval_bij_bij: fixes x::"'a::ordered_euclidean_space" assumes "\i. a$$i < b$$i \ u$$i < v$$i" - shows "interval_bij (a,b) (u,v) (interval_bij (u,v) (a,b) x) = x" - unfolding interval_bij_def split_conv euclidean_eq[where 'a='a] - apply(rule,insert assms,erule_tac x=i in allE) by auto +lemma interval_bij_bij: + "\(i::'a::ordered_euclidean_space)\Basis. a\i < b\i \ u\i < v\i \ + interval_bij (a,b) (u,v) (interval_bij (u,v) (a,b) x) = x" + by (auto simp: interval_bij_def euclidean_eq_iff[where 'a='a]) end diff -r 46be26e02456 -r 899c9c4e4a4c src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space.thy --- a/src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space.thy Fri Dec 14 14:46:01 2012 +0100 +++ b/src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space.thy Fri Dec 14 15:46:01 2012 +0100 @@ -328,283 +328,18 @@ fixes f:: "'a \ real ^'n" assumes fP: "finite P" and fPs: "\Q. Q \ P \ norm (setsum f Q) \ e" shows "setsum (\x. norm (f x)) P \ 2 * real CARD('n) * e" -proof - - let ?d = "real CARD('n)" - let ?nf = "\x. norm (f x)" - let ?U = "UNIV :: 'n set" - have th0: "setsum (\x. setsum (\i. \f x $ i\) ?U) P = setsum (\i. setsum (\x. \f x $ i\) P) ?U" - by (rule setsum_commute) - have th1: "2 * ?d * e = of_nat (card ?U) * (2 * e)" by (simp add: real_of_nat_def) - have "setsum ?nf P \ setsum (\x. setsum (\i. \f x $ i\) ?U) P" - apply (rule setsum_mono) - apply (rule norm_le_l1_cart) - done - also have "\ \ 2 * ?d * e" - unfolding th0 th1 - proof(rule setsum_bounded) - fix i assume i: "i \ ?U" - let ?Pp = "{x. x\ P \ f x $ i \ 0}" - let ?Pn = "{x. x \ P \ f x $ i < 0}" - have thp: "P = ?Pp \ ?Pn" by auto - have thp0: "?Pp \ ?Pn ={}" by auto - have PpP: "?Pp \ P" and PnP: "?Pn \ P" by blast+ - have Ppe:"setsum (\x. \f x $ i\) ?Pp \ e" - using component_le_norm_cart[of "setsum (\x. f x) ?Pp" i] fPs[OF PpP] - by (auto intro: abs_le_D1) - have Pne: "setsum (\x. \f x $ i\) ?Pn \ e" - using component_le_norm_cart[of "setsum (\x. - f x) ?Pn" i] fPs[OF PnP] - by (auto simp add: setsum_negf intro: abs_le_D1) - have "setsum (\x. \f x $ i\) P = setsum (\x. \f x $ i\) ?Pp + setsum (\x. \f x $ i\) ?Pn" - apply (subst thp) - apply (rule setsum_Un_zero) - using fP thp0 apply auto - done - also have "\ \ 2*e" using Pne Ppe by arith - finally show "setsum (\x. \f x $ i\) P \ 2*e" . - qed - finally show ?thesis . -qed - -lemma if_distr: "(if P then f else g) $ i = (if P then f $ i else g $ i)" by simp - -lemma split_dimensions'[consumes 1]: - assumes "k < DIM('a::euclidean_space^'b)" - obtains i j where "i < CARD('b::finite)" - and "j < DIM('a::euclidean_space)" - and "k = j + i * DIM('a::euclidean_space)" - using split_times_into_modulo[OF assms[simplified]] . - -lemma cart_euclidean_bound[intro]: - assumes j:"j < DIM('a::euclidean_space)" - shows "j + \' (i::'b::finite) * DIM('a) < CARD('b) * DIM('a::euclidean_space)" - using linear_less_than_times[OF pi'_range j, of i] . - -lemma (in euclidean_space) forall_CARD_DIM: - "(\i (\(i::'b::finite) j. j P (j + \' i * DIM('a)))" - (is "?l \ ?r") -proof (safe elim!: split_times_into_modulo) - fix i :: 'b and j - assume "j < DIM('a)" - note linear_less_than_times[OF pi'_range[of i] this] - moreover assume "?l" - ultimately show "P (j + \' i * DIM('a))" by auto -next - fix i j - assume "i < CARD('b)" "j < DIM('a)" and "?r" - from `?r`[rule_format, OF `j < DIM('a)`, of "\ i"] `i < CARD('b)` - show "P (j + i * DIM('a))" by simp -qed - -lemma (in euclidean_space) exists_CARD_DIM: - "(\i (\i::'b::finite. \j' i * DIM('a)))" - using forall_CARD_DIM[where 'b='b, of "\x. \ P x"] by blast - -lemma forall_CARD: - "(\i (\i::'b::finite. P (\' i))" - using forall_CARD_DIM[where 'a=real, of P] by simp - -lemma exists_CARD: - "(\i (\i::'b::finite. P (\' i))" - using exists_CARD_DIM[where 'a=real, of P] by simp - -lemmas cart_simps = forall_CARD_DIM exists_CARD_DIM forall_CARD exists_CARD - -lemma cart_euclidean_nth[simp]: - fixes x :: "('a::euclidean_space, 'b::finite) vec" - assumes j:"j < DIM('a)" - shows "x $$ (j + \' i * DIM('a)) = x $ i $$ j" - unfolding euclidean_component_def inner_vec_def basis_eq_pi'[OF j] if_distrib cond_application_beta - by (simp add: setsum_cases) - -lemma real_euclidean_nth: - fixes x :: "real^'n" - shows "x $$ \' i = (x $ i :: real)" - using cart_euclidean_nth[where 'a=real, of 0 x i] by simp - -lemmas nth_conv_component = real_euclidean_nth[symmetric] - -lemma mult_split_eq: - fixes A :: nat assumes "x < A" "y < A" - shows "x + i * A = y + j * A \ x = y \ i = j" -proof - assume *: "x + i * A = y + j * A" - { fix x y i j assume "i < j" "x < A" and *: "x + i * A = y + j * A" - hence "x + i * A < Suc i * A" using `x < A` by simp - also have "\ \ j * A" using `i < j` unfolding mult_le_cancel2 by simp - also have "\ \ y + j * A" by simp - finally have "i = j" using * by simp } - note eq = this - - have "i = j" - proof (cases rule: linorder_cases) - assume "i < j" - from eq[OF this `x < A` *] show "i = j" by simp - next - assume "j < i" - from eq[OF this `y < A` *[symmetric]] show "i = j" by simp - qed simp - thus "x = y \ i = j" using * by simp -qed simp + using setsum_norm_allsubsets_bound[OF assms] + by (simp add: DIM_cart Basis_real_def) instance vec :: (ordered_euclidean_space, finite) ordered_euclidean_space proof fix x y::"'a^'b" - show "(x \ y) = (\i y $$ i)" - unfolding less_eq_vec_def apply(subst eucl_le) by (simp add: cart_simps) - show"(x < y) = (\i i. if i = k then 1 else 0)" - -lemma basis_component [simp]: "cart_basis k $ i = (if k=i then 1 else 0)" - unfolding cart_basis_def by simp - -lemma norm_basis[simp]: - shows "norm (cart_basis k :: real ^'n) = 1" - apply (simp add: cart_basis_def norm_eq_sqrt_inner) unfolding inner_vec_def - apply (vector delta_mult_idempotent) - using setsum_delta[of "UNIV :: 'n set" "k" "\k. 1::real"] apply auto - done - -lemma norm_basis_1: "norm(cart_basis 1 :: real ^'n::{finite,one}) = 1" - by (rule norm_basis) - -lemma vector_choose_size: "0 <= c ==> \(x::real^'n). norm x = c" - by (rule exI[where x="c *\<^sub>R cart_basis arbitrary"]) simp - -lemma vector_choose_dist: - assumes e: "0 <= e" - shows "\(y::real^'n). dist x y = e" -proof - - from vector_choose_size[OF e] obtain c:: "real ^'n" where "norm c = e" - by blast - then have "dist x (x - c) = e" by (simp add: dist_norm) - then show ?thesis by blast + show "(x \ y) = (\i\Basis. x \ i \ y \ i)" + unfolding less_eq_vec_def apply(subst eucl_le) by (simp add: Basis_vec_def inner_axis) + show"(x < y) = (\i\Basis. x \ i < y \ i)" + unfolding less_vec_def apply(subst eucl_less) by (simp add: Basis_vec_def inner_axis) qed -lemma basis_inj[intro]: "inj (cart_basis :: 'n \ real ^'n)" - by (simp add: inj_on_def vec_eq_iff) - -lemma basis_expansion: "setsum (\i. (x$i) *s cart_basis i) UNIV = (x::('a::ring_1) ^'n)" - (is "?lhs = ?rhs" is "setsum ?f ?S = _") - by (auto simp add: vec_eq_iff - if_distrib setsum_delta[of "?S", where ?'b = "'a", simplified] cong del: if_weak_cong) - -lemma smult_conv_scaleR: "c *s x = scaleR c x" - unfolding vector_scalar_mult_def scaleR_vec_def by simp - -lemma basis_expansion': "setsum (\i. (x$i) *\<^sub>R cart_basis i) UNIV = x" - by (rule basis_expansion [where 'a=real, unfolded smult_conv_scaleR]) - -lemma basis_expansion_unique: - "setsum (\i. f i *s cart_basis i) UNIV = (x::('a::comm_ring_1) ^'n) \ (\i. f i = x$i)" - by (simp add: vec_eq_iff setsum_delta if_distrib cong del: if_weak_cong) - -lemma dot_basis: "cart_basis i \ x = x$i" "x \ (cart_basis i) = (x$i)" - by (auto simp add: inner_vec_def cart_basis_def cond_application_beta if_distrib setsum_delta - cong del: if_weak_cong) - -lemma inner_basis: - fixes x :: "'a::{real_inner, real_algebra_1} ^ 'n" - shows "inner (cart_basis i) x = inner 1 (x $ i)" - and "inner x (cart_basis i) = inner (x $ i) 1" - unfolding inner_vec_def cart_basis_def - by (auto simp add: cond_application_beta if_distrib setsum_delta cong del: if_weak_cong) - -lemma basis_eq_0: "cart_basis i = (0::'a::semiring_1^'n) \ False" - by (auto simp add: vec_eq_iff) - -lemma basis_nonzero: "cart_basis k \ (0:: 'a::semiring_1 ^'n)" - by (simp add: basis_eq_0) - -text {* some lemmas to map between Eucl and Cart *} -lemma basis_real_n[simp]:"(basis (\' i)::real^'a) = cart_basis i" - unfolding basis_vec_def using pi'_range[where 'n='a] - by (auto simp: vec_eq_iff axis_def) - -subsection {* Orthogonality on cartesian products *} - -lemma orthogonal_basis: "orthogonal (cart_basis i) x \ x$i = (0::real)" - by (auto simp add: orthogonal_def inner_vec_def cart_basis_def if_distrib - cond_application_beta setsum_delta cong del: if_weak_cong) - -lemma orthogonal_basis_basis: "orthogonal (cart_basis i :: real^'n) (cart_basis j) \ i \ j" - unfolding orthogonal_basis[of i] basis_component[of j] by simp - -subsection {* Linearity on cartesian products *} - -lemma linear_vmul_component: - assumes "linear f" - shows "linear (\x. f x $ k *\<^sub>R v)" - using assms by (auto simp add: linear_def algebra_simps) - - -subsection {* Adjoints on cartesian products *} - -text {* TODO: The following lemmas about adjoints should hold for any -Hilbert space (i.e. complete inner product space). -(see \url{http://en.wikipedia.org/wiki/Hermitian_adjoint}) -*} - -lemma adjoint_works_lemma: - fixes f:: "real ^'n \ real ^'m" - assumes lf: "linear f" - shows "\x y. f x \ y = x \ adjoint f y" -proof - - let ?N = "UNIV :: 'n set" - let ?M = "UNIV :: 'm set" - have fN: "finite ?N" by simp - have fM: "finite ?M" by simp - { fix y:: "real ^ 'm" - let ?w = "(\ i. (f (cart_basis i) \ y)) :: real ^ 'n" - { fix x - have "f x \ y = f (setsum (\i. (x$i) *\<^sub>R cart_basis i) ?N) \ y" - by (simp only: basis_expansion') - also have "\ = (setsum (\i. (x$i) *\<^sub>R f (cart_basis i)) ?N) \ y" - unfolding linear_setsum[OF lf fN] - by (simp add: linear_cmul[OF lf]) - finally have "f x \ y = x \ ?w" - by (simp add: inner_vec_def setsum_left_distrib - setsum_right_distrib setsum_commute[of _ ?M ?N] field_simps) - } - } - then show ?thesis - unfolding adjoint_def some_eq_ex[of "\f'. \x y. f x \ y = x \ f' y"] - using choice_iff[of "\a b. \x. f x \ a = x \ b "] - by metis -qed - -lemma adjoint_works: - fixes f:: "real ^'n \ real ^'m" - assumes lf: "linear f" - shows "x \ adjoint f y = f x \ y" - using adjoint_works_lemma[OF lf] by metis - -lemma adjoint_linear: - fixes f:: "real ^'n \ real ^'m" - assumes lf: "linear f" - shows "linear (adjoint f)" - unfolding linear_def vector_eq_ldot[where 'a="real^'n", symmetric] apply safe - unfolding inner_simps smult_conv_scaleR adjoint_works[OF lf] by auto - -lemma adjoint_clauses: - fixes f:: "real ^'n \ real ^'m" - assumes lf: "linear f" - shows "x \ adjoint f y = f x \ y" - and "adjoint f y \ x = y \ f x" - by (simp_all add: adjoint_works[OF lf] inner_commute) - -lemma adjoint_adjoint: - fixes f:: "real ^'n \ real ^'m" - assumes lf: "linear f" - shows "adjoint (adjoint f) = f" - by (rule adjoint_unique, simp add: adjoint_clauses [OF lf]) - - subsection {* Matrix operations *} text{* Matrix notation. NB: an MxN matrix is of type @{typ "'a^'n^'m"}, not @{typ "'a^'m^'n"} *} @@ -680,10 +415,10 @@ apply auto apply (subst vec_eq_iff) apply clarify - apply (clarsimp simp add: matrix_vector_mult_def cart_basis_def if_distrib cond_application_beta vec_eq_iff cong del: if_weak_cong) - apply (erule_tac x="cart_basis ia" in allE) + apply (clarsimp simp add: matrix_vector_mult_def if_distrib cond_application_beta vec_eq_iff cong del: if_weak_cong) + apply (erule_tac x="axis ia 1" in allE) apply (erule_tac x="i" in allE) - apply (auto simp add: cart_basis_def if_distrib cond_application_beta + apply (auto simp add: if_distrib cond_application_beta axis_def setsum_delta[OF finite] cong del: if_weak_cong) done @@ -728,25 +463,27 @@ by (simp add: matrix_vector_mult_def vec_eq_iff column_def mult_commute) lemma vector_componentwise: - "(x::'a::ring_1^'n) = (\ j. setsum (\i. (x$i) * (cart_basis i :: 'a^'n)$j) (UNIV :: 'n set))" - apply (subst basis_expansion[symmetric]) - apply (vector vec_eq_iff setsum_component) - done + "(x::'a::ring_1^'n) = (\ j. \i\UNIV. (x$i) * (axis i 1 :: 'a^'n) $ j)" + by (simp add: axis_def if_distrib setsum_cases vec_eq_iff) + +lemma basis_expansion: "setsum (\i. (x$i) *s axis i 1) UNIV = (x::('a::ring_1) ^'n)" + by (auto simp add: axis_def vec_eq_iff if_distrib setsum_cases cong del: if_weak_cong) lemma linear_componentwise: fixes f:: "real ^'m \ real ^ _" assumes lf: "linear f" - shows "(f x)$j = setsum (\i. (x$i) * (f (cart_basis i)$j)) (UNIV :: 'm set)" (is "?lhs = ?rhs") + shows "(f x)$j = setsum (\i. (x$i) * (f (axis i 1)$j)) (UNIV :: 'm set)" (is "?lhs = ?rhs") proof - let ?M = "(UNIV :: 'm set)" let ?N = "(UNIV :: 'n set)" have fM: "finite ?M" by simp - have "?rhs = (setsum (\i.(x$i) *\<^sub>R f (cart_basis i) ) ?M)$j" - unfolding vector_smult_component[symmetric] smult_conv_scaleR - unfolding setsum_component[of "(\i.(x$i) *\<^sub>R f (cart_basis i :: real^'m))" ?M] - .. + have "?rhs = (setsum (\i.(x$i) *\<^sub>R f (axis i 1) ) ?M)$j" + unfolding setsum_component by simp then show ?thesis - unfolding linear_setsum_mul[OF lf fM, symmetric] basis_expansion' .. + unfolding linear_setsum_mul[OF lf fM, symmetric] + unfolding scalar_mult_eq_scaleR[symmetric] + unfolding basis_expansion + by simp qed text{* Inverse matrices (not necessarily square) *} @@ -761,7 +498,7 @@ text{* Correspondence between matrices and linear operators. *} definition matrix :: "('a::{plus,times, one, zero}^'m \ 'a ^ 'n) \ 'a^'m^'n" - where "matrix f = (\ i j. (f(cart_basis j))$i)" + where "matrix f = (\ i j. (f(axis j 1))$i)" lemma matrix_vector_mul_linear: "linear(\x. A *v (x::real ^ _))" by (simp add: linear_def matrix_vector_mult_def vec_eq_iff @@ -831,103 +568,8 @@ ultimately show ?thesis by metis qed - -subsection {* Standard bases are a spanning set, and obviously finite. *} - -lemma span_stdbasis:"span {cart_basis i :: real^'n | i. i \ (UNIV :: 'n set)} = UNIV" - apply (rule set_eqI) - apply auto - apply (subst basis_expansion'[symmetric]) - apply (rule span_setsum) - apply simp - apply auto - apply (rule span_mul) - apply (rule span_superset) - apply auto - done - -lemma finite_stdbasis: "finite {cart_basis i ::real^'n |i. i\ (UNIV:: 'n set)}" (is "finite ?S") -proof - - have "?S = cart_basis ` UNIV" by blast - then show ?thesis by auto -qed - -lemma card_stdbasis: "card {cart_basis i ::real^'n |i. i\ (UNIV :: 'n set)} = CARD('n)" (is "card ?S = _") -proof - - have "?S = cart_basis ` UNIV" by blast - then show ?thesis using card_image[OF basis_inj] by simp -qed - -lemma independent_stdbasis_lemma: - assumes x: "(x::real ^ 'n) \ span (cart_basis ` S)" - and iS: "i \ S" - shows "(x$i) = 0" -proof - - let ?U = "UNIV :: 'n set" - let ?B = "cart_basis ` S" - let ?P = "{(x::real^_). \i\ ?U. i \ S \ x$i =0}" - { fix x::"real^_" assume xS: "x\ ?B" - from xS have "x \ ?P" by auto } - moreover - have "subspace ?P" - by (auto simp add: subspace_def) - ultimately show ?thesis - using x span_induct[of x ?B ?P] iS by blast -qed - -lemma independent_stdbasis: "independent {cart_basis i ::real^'n |i. i\ (UNIV :: 'n set)}" -proof - - let ?I = "UNIV :: 'n set" - let ?b = "cart_basis :: _ \ real ^'n" - let ?B = "?b ` ?I" - have eq: "{?b i|i. i \ ?I} = ?B" by auto - { assume d: "dependent ?B" - then obtain k where k: "k \ ?I" "?b k \ span (?B - {?b k})" - unfolding dependent_def by auto - have eq1: "?B - {?b k} = ?B - ?b ` {k}" by simp - have eq2: "?B - {?b k} = ?b ` (?I - {k})" - unfolding eq1 - apply (rule inj_on_image_set_diff[symmetric]) - apply (rule basis_inj) using k(1) - apply auto - done - from k(2) have th0: "?b k \ span (?b ` (?I - {k}))" unfolding eq2 . - from independent_stdbasis_lemma[OF th0, of k, simplified] - have False by simp } - then show ?thesis unfolding eq dependent_def .. -qed - lemma vector_sub_project_orthogonal_cart: "(b::real^'n) \ (x - ((b \ x) / (b \ b)) *s b) = 0" - unfolding inner_simps smult_conv_scaleR by auto - -lemma linear_eq_stdbasis_cart: - assumes lf: "linear (f::real^'m \ _)" and lg: "linear g" - and fg: "\i. f (cart_basis i) = g(cart_basis i)" - shows "f = g" -proof - - let ?U = "UNIV :: 'm set" - let ?I = "{cart_basis i:: real^'m|i. i \ ?U}" - { fix x assume x: "x \ (UNIV :: (real^'m) set)" - from equalityD2[OF span_stdbasis] - have IU: " (UNIV :: (real^'m) set) \ span ?I" by blast - from linear_eq[OF lf lg IU] fg x - have "f x = g x" unfolding Ball_def mem_Collect_eq by metis - } - then show ?thesis by auto -qed - -lemma bilinear_eq_stdbasis_cart: - assumes bf: "bilinear (f:: real^'m \ real^'n \ _)" - and bg: "bilinear g" - and fg: "\i j. f (cart_basis i) (cart_basis j) = g (cart_basis i) (cart_basis j)" - shows "f = g" -proof - - from fg have th: "\x \ {cart_basis i| i. i\ (UNIV :: 'm set)}. - \y\ {cart_basis j |j. j \ (UNIV :: 'n set)}. f x y = g x y" - by blast - from bilinear_eq[OF bf bg equalityD2[OF span_stdbasis] equalityD2[OF span_stdbasis] th] - show ?thesis by blast -qed + unfolding inner_simps scalar_mult_eq_scaleR by auto lemma left_invertible_transpose: "(\(B). B ** transpose (A) = mat (1::'a::comm_semiring_1)) \ (\(B). A ** B = mat 1)" @@ -1043,7 +685,7 @@ unfolding y[symmetric] apply (rule span_setsum[OF fU]) apply clarify - unfolding smult_conv_scaleR + unfolding scalar_mult_eq_scaleR apply (rule span_mul) apply (rule span_superset) unfolding columns_def @@ -1056,7 +698,7 @@ let ?P = "\(y::real ^'n). \(x::real^'m). setsum (\i. (x$i) *s column i A) ?U = y" { fix y have "?P y" - proof (rule span_induct_alt[of ?P "columns A", folded smult_conv_scaleR]) + proof (rule span_induct_alt[of ?P "columns A", folded scalar_mult_eq_scaleR]) show "\x\real ^ 'm. setsum (\i. (x$i) *s column i A) ?U = 0" by (rule exI[where x=0], simp) next @@ -1159,25 +801,12 @@ dot_rowvector_columnvector matrix_transpose_mul matrix_mul_assoc .. -lemma infnorm_cart:"infnorm (x::real^'n) = Sup {abs(x$i) |i. i\ (UNIV :: 'n set)}" - unfolding infnorm_def apply(rule arg_cong[where f=Sup]) apply safe - apply(rule_tac x="\ i" in exI) defer - apply(rule_tac x="\' i" in exI) - unfolding nth_conv_component - using pi'_range apply auto - done - -lemma infnorm_set_image_cart: "{abs(x$i) |i. i\ (UNIV :: _ set)} = - (\i. abs(x$i)) ` (UNIV)" by blast - -lemma infnorm_set_lemma_cart: - "finite {abs((x::'a::abs ^'n)$i) |i. i\ (UNIV :: 'n set)}" - "{abs(x$i) |i. i\ (UNIV :: 'n::finite set)} \ {}" - unfolding infnorm_set_image_cart by auto +lemma infnorm_cart:"infnorm (x::real^'n) = Sup {abs(x$i) |i. i\UNIV}" + by (simp add: infnorm_def inner_axis Basis_vec_def) (metis (lifting) inner_axis real_inner_1_right) lemma component_le_infnorm_cart: "\x$i\ \ infnorm (x::real^'n)" - unfolding nth_conv_component - using component_le_infnorm[of x] . + using Basis_le_infnorm[of "axis i 1" x] + by (simp add: Basis_vec_def axis_eq_axis inner_axis) lemma continuous_component: "continuous F f \ continuous F (\x. f x $ i)" unfolding continuous_def by (rule tendsto_vec_nth) @@ -1371,7 +1000,7 @@ and "{c .. d} \ {a<.. (\i. c$i \ d$i) --> (\i. a$i < c$i \ d$i < b$i)" (is ?th2) and "{c<.. {a .. b} \ (\i. c$i < d$i) --> (\i. a$i \ c$i \ d$i \ b$i)" (is ?th3) and "{c<.. {a<.. (\i. c$i < d$i) --> (\i. a$i \ c$i \ d$i \ b$i)" (is ?th4) - using subset_interval[of c d a b] by (simp_all add: cart_simps real_euclidean_nth) + using subset_interval[of c d a b] by (simp_all add: Basis_vec_def inner_axis) lemma disjoint_interval_cart: fixes a::"real^'n" @@ -1379,7 +1008,7 @@ and "{a .. b} \ {c<.. (\i. (b$i < a$i \ d$i \ c$i \ b$i \ c$i \ d$i \ a$i))" (is ?th2) and "{a<.. {c .. d} = {} \ (\i. (b$i \ a$i \ d$i < c$i \ b$i \ c$i \ d$i \ a$i))" (is ?th3) and "{a<.. {c<.. (\i. (b$i \ a$i \ d$i \ c$i \ b$i \ c$i \ d$i \ a$i))" (is ?th4) - using disjoint_interval[of a b c d] by (simp_all add: cart_simps real_euclidean_nth) + using disjoint_interval[of a b c d] by (simp_all add: Basis_vec_def inner_axis) lemma inter_interval_cart: fixes a :: "'a::linorder^'n" @@ -1400,7 +1029,7 @@ 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)" - by (simp add: is_interval_def Ball_def cart_simps real_euclidean_nth) + by (simp add: is_interval_def Ball_def Basis_vec_def inner_axis imp_ex) lemma closed_halfspace_component_le_cart: "closed {x::real^'n. x$i \ a}" by (simp add: closed_Collect_le) @@ -1416,27 +1045,15 @@ lemma Lim_component_le_cart: fixes f :: "'a \ real^'n" - assumes "(f ---> l) net" "\ (trivial_limit net)" "eventually (\x. f(x)$i \ b) net" + assumes "(f ---> l) net" "\ (trivial_limit net)" "eventually (\x. f x $i \ b) net" shows "l$i \ b" -proof - - { fix x - have "x \ {x::real^'n. inner (cart_basis i) x \ b} \ x$i \ b" - unfolding inner_basis by auto } - then show ?thesis using Lim_in_closed_set[of "{x. inner (cart_basis i) x \ b}" f net l] - using closed_halfspace_le[of "(cart_basis i)::real^'n" b] and assms(1,2,3) by auto -qed + by (rule tendsto_le[OF assms(2) tendsto_const tendsto_vec_nth, OF assms(1, 3)]) lemma Lim_component_ge_cart: fixes f :: "'a \ real^'n" assumes "(f ---> l) net" "\ (trivial_limit net)" "eventually (\x. b \ (f x)$i) net" shows "b \ l$i" -proof - - { fix x - have "x \ {x::real^'n. inner (cart_basis i) x \ b} \ x$i \ b" - unfolding inner_basis by auto } - then show ?thesis using Lim_in_closed_set[of "{x. inner (cart_basis i) x \ b}" f net l] - using closed_halfspace_ge[of b "(cart_basis i)::real^'n"] and assms(1,2,3) by auto -qed + by (rule tendsto_le[OF assms(2) tendsto_vec_nth tendsto_const, OF assms(1, 3)]) lemma Lim_component_eq_cart: fixes f :: "'a \ real^'n" @@ -1449,8 +1066,8 @@ lemma connected_ivt_component_cart: fixes x :: "real^'n" shows "connected s \ x \ s \ y \ s \ x$k \ a \ a \ y$k \ (\z\s. z$k = a)" - using connected_ivt_hyperplane[of s x y "(cart_basis k)::real^'n" a] - by (auto simp add: inner_basis) + using connected_ivt_hyperplane[of s x y "axis k 1" a] + by (auto simp add: inner_axis inner_commute) lemma subspace_substandard_cart: "subspace {x::real^_. (\i. P i \ x$i = 0)}" unfolding subspace_def by auto @@ -1468,20 +1085,14 @@ lemma dim_substandard_cart: "dim {x::real^'n. \i. i \ d \ x$i = 0} = card d" (is "dim ?A = _") proof - - have *: "{x. \i \' ` d \ x $$ i = 0} = - {x::real^'n. \i. i \ d \ x$i = 0}" - apply safe - apply (erule_tac x="\' i" in allE) defer - apply (erule_tac x="\ i" in allE) - unfolding image_iff real_euclidean_nth[symmetric] - apply (auto simp: pi'_inj[THEN inj_eq]) - done - have " \' ` d \ {..i\Basis. i \ ?a ` d \ x \ i = 0} = ?A" + by (auto simp: image_iff Basis_vec_def axis_eq_axis inner_axis) + have "?a ` d \ Basis" + by (auto simp: Basis_vec_def) thus ?thesis - using dim_substandard[of "\' ` d", where 'a="real^'n"] - unfolding * using card_image[of "\'" d] using pi'_inj unfolding inj_on_def - by auto + using dim_substandard[of "?a ` d"] card_image[of ?a d] + by (auto simp: axis_eq_axis inj_on_def *) qed lemma affinity_inverses: @@ -1513,32 +1124,24 @@ using vector_affinity_eq[where m=m and x=x and y=y and c=c] by metis -lemma const_vector_cart:"((\ i. d)::real^'n) = (\\ i. d)" - apply(subst euclidean_eq) -proof safe - case goal1 - hence *: "(basis i::real^'n) = cart_basis (\ i)" - unfolding basis_real_n[symmetric] by auto - have "((\ i. d)::real^'n) $$ i = d" unfolding euclidean_component_def * - unfolding dot_basis by auto - thus ?case using goal1 by auto -qed - +lemma vector_cart: + fixes f :: "real^'n \ real" + shows "(\ i. f (axis i 1)) = (\i\Basis. f i *\<^sub>R i)" + unfolding euclidean_eq_iff[where 'a="real^'n"] + by simp (simp add: Basis_vec_def inner_axis) + +lemma const_vector_cart:"((\ i. d)::real^'n) = (\i\Basis. d *\<^sub>R i)" + by (rule vector_cart) subsection "Convex Euclidean Space" -lemma Cart_1:"(1::real^'n) = (\\ i. 1)" - apply(subst euclidean_eq) -proof safe - case goal1 - thus ?case - using nth_conv_component[THEN sym,where i1="\ i" and x1="1::real^'n"] by auto -qed +lemma Cart_1:"(1::real^'n) = \Basis" + using const_vector_cart[of 1] by (simp add: one_vec_def) declare vector_add_ldistrib[simp] vector_ssub_ldistrib[simp] vector_smult_assoc[simp] vector_smult_rneg[simp] declare vector_sadd_rdistrib[simp] vector_sub_rdistrib[simp] -lemmas vector_component_simps = vector_minus_component vector_smult_component vector_add_component less_eq_vec_def vec_lambda_beta basis_component vector_uminus_component +lemmas vector_component_simps = vector_minus_component vector_smult_component vector_add_component less_eq_vec_def vec_lambda_beta vector_uminus_component lemma convex_box_cart: assumes "\i. convex {x. P i x}" @@ -1551,95 +1154,20 @@ lemma unit_interval_convex_hull_cart: "{0::real^'n .. 1} = convex hull {x. \i. (x$i = 0) \ (x$i = 1)}" (is "?int = convex hull ?points") unfolding Cart_1 unit_interval_convex_hull[where 'a="real^'n"] - apply(rule arg_cong[where f="\x. convex hull x"]) apply(rule set_eqI) unfolding mem_Collect_eq - apply safe apply(erule_tac x="\' i" in allE) unfolding nth_conv_component defer - apply(erule_tac x="\ i" in allE) - apply auto - done + by (rule arg_cong[where f="\x. convex hull x"]) (simp add: Basis_vec_def inner_axis) lemma cube_convex_hull_cart: assumes "0 < d" obtains s::"(real^'n) set" where "finite s" "{x - (\ i. d) .. x + (\ i. d)} = convex hull s" proof - - obtain s where s: "finite s" "{x - (\\ i. d)..x + (\\ i. d)} = convex hull s" - by (rule cube_convex_hull [OF assms]) - show thesis - apply(rule that[OF s(1)]) unfolding s(2)[symmetric] const_vector_cart .. + from cube_convex_hull [OF assms, of x] guess s . + with that[of s] show thesis by (simp add: const_vector_cart) qed -lemma std_simplex_cart: - "(insert (0::real^'n) { cart_basis i | i. i\UNIV}) = - (insert 0 { basis i | i. i i. u$i + (x$i - a$i) / (b$i - a$i) * (v$i - u$i))::real^'n)" - unfolding interval_bij_def apply(rule ext)+ apply safe - unfolding vec_eq_iff vec_lambda_beta unfolding nth_conv_component - apply rule - apply (subst euclidean_lambda_beta) - using pi'_range apply auto - done - -lemma interval_bij_affine_cart: - "interval_bij (a,b) (u,v) = (\x. (\ i. (v$i - u$i) / (b$i - a$i) * x$i) + - (\ i. u$i - (v$i - u$i) / (b$i - a$i) * a$i)::real^'n)" - apply rule - unfolding vec_eq_iff interval_bij_cart vector_component_simps - apply (auto simp add: field_simps add_divide_distrib[symmetric]) - done - subsection "Derivative" -lemma has_derivative_vmul_component_cart: - fixes c :: "real^'a \ real^'b" and v :: "real^'c" - assumes "(c has_derivative c') net" - shows "((\x. c(x)$k *\<^sub>R v) has_derivative (\x. (c' x)$k *\<^sub>R v)) net" - unfolding nth_conv_component by (intro has_derivative_intros assms) - lemma differentiable_at_imp_differentiable_on: "(\x\(s::(real^'n) set). f differentiable at x) \ f differentiable_on s" unfolding differentiable_on_def by(auto intro!: differentiable_at_withinI) @@ -1662,56 +1190,14 @@ subsection {* Component of the differential must be zero if it exists at a local maximum or minimum for that corresponding component. *} -lemma differential_zero_maxmin_component: +lemma differential_zero_maxmin_cart: fixes f::"real^'a \ real^'b" assumes "0 < e" "((\y \ ball x e. (f y)$k \ (f x)$k) \ (\y\ball x e. (f x)$k \ (f y)$k))" - "f differentiable (at x)" shows "jacobian f (at x) $ k = 0" -(* FIXME: reuse proof of generic differential_zero_maxmin_component*) -proof (rule ccontr) - def D \ "jacobian f (at x)" - assume "jacobian f (at x) $ k \ 0" - then obtain j where j:"D$k$j \ 0" unfolding vec_eq_iff D_def by auto - hence *: "abs (jacobian f (at x) $ k $ j) / 2 > 0" - unfolding D_def by auto - note as = assms(3)[unfolded jacobian_works has_derivative_at_alt] - guess e' using as[THEN conjunct2,rule_format,OF *] .. note e' = this - guess d using real_lbound_gt_zero[OF assms(1) e'[THEN conjunct1]] .. note d = this - { fix c - assume "abs c \ d" - hence *:"norm (x + c *\<^sub>R cart_basis j - x) < e'" - using norm_basis[of j] d by auto - have "\(f (x + c *\<^sub>R cart_basis j) - f x - D *v (c *\<^sub>R cart_basis j)) $ k\ \ - norm (f (x + c *\<^sub>R cart_basis j) - f x - D *v (c *\<^sub>R cart_basis j))" - by (rule component_le_norm_cart) - also have "\ \ \D $ k $ j\ / 2 * \c\" - using e'[THEN conjunct2,rule_format,OF *] and norm_basis[of j] - unfolding D_def[symmetric] by auto - finally - have "\(f (x + c *\<^sub>R cart_basis j) - f x - D *v (c *\<^sub>R cart_basis j)) $ k\ \ - \D $ k $ j\ / 2 * \c\" by simp - hence "\f (x + c *\<^sub>R cart_basis j) $ k - f x $ k - c * D $ k $ j\ \ - \D $ k $ j\ / 2 * \c\" - unfolding vector_component_simps matrix_vector_mul_component - unfolding smult_conv_scaleR[symmetric] - unfolding inner_simps dot_basis smult_conv_scaleR by simp - } note * = this - have "x + d *\<^sub>R cart_basis j \ ball x e" "x - d *\<^sub>R cart_basis j \ ball x e" - unfolding mem_ball dist_norm using norm_basis[of j] d by auto - hence **: "((f (x - d *\<^sub>R cart_basis j))$k \ (f x)$k \ (f (x + d *\<^sub>R cart_basis j))$k \ (f x)$k) \ - ((f (x - d *\<^sub>R cart_basis j))$k \ (f x)$k \ (f (x + d *\<^sub>R cart_basis j))$k \ (f x)$k)" - using assms(2) by auto - have ***: "\y y1 y2 d dx::real. (y1\y\y2\y) \ (y\y1\y\y2) \ - d < abs dx \ abs(y1 - y - - dx) \ d \ (abs (y2 - y - dx) \ d) \ False" by arith - show False - apply (rule ***[OF **, where dx="d * D $ k $ j" and d="\D $ k $ j\ / 2 * \d\"]) - using *[of "-d"] and *[of d] and d[THEN conjunct1] and j - unfolding mult_minus_left - unfolding abs_mult diff_minus_eq_add scaleR_minus_left - unfolding algebra_simps - apply (auto intro: mult_pos_pos) - done -qed - + "f differentiable (at x)" + shows "jacobian f (at x) $ k = 0" + using differential_zero_maxmin_component[of "axis k 1" e x f] assms + vector_cart[of "\j. frechet_derivative f (at x) j $ k"] + by (simp add: Basis_vec_def axis_eq_axis inner_axis jacobian_def matrix_def) subsection {* Lemmas for working on @{typ "real^1"} *} @@ -1775,25 +1261,6 @@ end -(* "lift" from 'a to 'a^1 and "drop" from 'a^1 to 'a -- FIXME: potential use of transfer *) - -abbreviation vec1:: "'a \ 'a ^ 1" where "vec1 x \ vec x" - -abbreviation dest_vec1:: "'a ^1 \ 'a" where "dest_vec1 x \ (x$1)" - -lemma vec1_dest_vec1[simp]: "vec1(dest_vec1 x) = x" - by (simp add: vec_eq_iff) - -lemma forall_vec1: "(\x. P x) \ (\x. P (vec1 x))" - by (metis vec1_dest_vec1(1)) - -lemma exists_vec1: "(\x. P x) \ (\x. P(vec1 x))" - by (metis vec1_dest_vec1(1)) - -lemma dest_vec1_eq[simp]: "dest_vec1 x = dest_vec1 y \ x = y" - by (metis vec1_dest_vec1(1)) - - subsection{* The collapse of the general concepts to dimension one. *} lemma vector_one: "(x::'a ^1) = (\ i. (x$1))" @@ -1863,438 +1330,10 @@ apply (simp add: forall_3) done -lemma range_vec1[simp]:"range vec1 = UNIV" - apply (rule set_eqI,rule) unfolding image_iff defer - apply (rule_tac x="dest_vec1 x" in bexI) - apply auto - done - -lemma dest_vec1_lambda: "dest_vec1(\ i. x i) = x 1" - by simp - -lemma dest_vec1_vec: "dest_vec1(vec x) = x" - by simp - -lemma dest_vec1_sum: assumes fS: "finite S" - shows "dest_vec1(setsum f S) = setsum (dest_vec1 o f) S" - apply (induct rule: finite_induct[OF fS]) - apply simp - apply auto - done - -lemma norm_vec1 [simp]: "norm(vec1 x) = abs(x)" - by (simp add: vec_def norm_real) - -lemma dist_vec1: "dist(vec1 x) (vec1 y) = abs(x - y)" - by (simp only: dist_real vec_component) -lemma abs_dest_vec1: "norm x = \dest_vec1 x\" - by (metis vec1_dest_vec1(1) norm_vec1) - -lemmas vec1_dest_vec1_simps = - forall_vec1 vec_add[symmetric] dist_vec1 vec_sub[symmetric] vec1_dest_vec1 norm_vec1 vector_smult_component - vec_inj[where 'b=1] vec_cmul[symmetric] smult_conv_scaleR[symmetric] o_def dist_real_def real_norm_def - -lemma bounded_linear_vec1: "bounded_linear (vec1::real\real^1)" - unfolding bounded_linear_def additive_def bounded_linear_axioms_def - unfolding smult_conv_scaleR[symmetric] - unfolding vec1_dest_vec1_simps - apply (rule conjI) defer - apply (rule conjI) defer - apply (rule_tac x=1 in exI) - apply auto - done - -lemma linear_vmul_dest_vec1: - fixes f:: "real^_ \ real^1" - shows "linear f \ linear (\x. dest_vec1(f x) *s v)" - unfolding smult_conv_scaleR - by (rule linear_vmul_component) - -lemma linear_from_scalars: - assumes lf: "linear (f::real^1 \ real^_)" - shows "f = (\x. dest_vec1 x *s column 1 (matrix f))" - unfolding smult_conv_scaleR - apply (rule ext) - apply (subst matrix_works[OF lf, symmetric]) - apply (auto simp add: vec_eq_iff matrix_vector_mult_def column_def mult_commute) - done - -lemma linear_to_scalars: - assumes lf: "linear (f::real ^'n \ real^1)" - shows "f = (\x. vec1(row 1 (matrix f) \ x))" - apply (rule ext) - apply (subst matrix_works[OF lf, symmetric]) - apply (simp add: vec_eq_iff matrix_vector_mult_def row_def inner_vec_def mult_commute) - done - -lemma dest_vec1_eq_0: "dest_vec1 x = 0 \ x = 0" - by (simp add: dest_vec1_eq[symmetric]) - -lemma setsum_scalars: - assumes fS: "finite S" - shows "setsum f S = vec1 (setsum (dest_vec1 o f) S)" - unfolding vec_setsum[OF fS] by simp - -lemma dest_vec1_wlog_le: - "(\(x::'a::linorder ^ 1) y. P x y \ P y x) - \ (\x y. dest_vec1 x <= dest_vec1 y ==> P x y) \ P x y" - apply (cases "dest_vec1 x \ dest_vec1 y") - apply simp - apply (subgoal_tac "dest_vec1 y \ dest_vec1 x") - apply auto - done - -text{* Lifting and dropping *} - -lemma continuous_on_o_dest_vec1: - fixes f::"real \ 'a::real_normed_vector" - assumes "continuous_on {a..b::real} f" - shows "continuous_on {vec1 a..vec1 b} (f o dest_vec1)" - using assms unfolding continuous_on_iff apply safe - apply (erule_tac x="x$1" in ballE,erule_tac x=e in allE) - apply safe - apply (rule_tac x=d in exI) - apply safe - unfolding o_def dist_real_def dist_real - apply (erule_tac x="dest_vec1 x'" in ballE) - apply (auto simp add:less_eq_vec_def) - done - -lemma continuous_on_o_vec1: - fixes f::"real^1 \ 'a::real_normed_vector" - assumes "continuous_on {a..b} f" - shows "continuous_on {dest_vec1 a..dest_vec1 b} (f o vec1)" - using assms unfolding continuous_on_iff - apply safe - apply (erule_tac x="vec x" in ballE,erule_tac x=e in allE) - apply safe - apply (rule_tac x=d in exI) - apply safe - unfolding o_def dist_real_def dist_real - apply (erule_tac x="vec1 x'" in ballE) - apply (auto simp add:less_eq_vec_def) - done - -lemma continuous_on_vec1:"continuous_on A (vec1::real\real^1)" - by (rule linear_continuous_on[OF bounded_linear_vec1]) - -lemma mem_interval_1: - fixes x :: "real^1" - shows "(x \ {a .. b} \ dest_vec1 a \ dest_vec1 x \ dest_vec1 x \ dest_vec1 b)" - and "(x \ {a<.. dest_vec1 a < dest_vec1 x \ dest_vec1 x < dest_vec1 b)" - by (simp_all add: vec_eq_iff less_vec_def less_eq_vec_def) - -lemma vec1_interval: - fixes a::"real" - shows "vec1 ` {a .. b} = {vec1 a .. vec1 b}" - and "vec1 ` {a<.. {a .. b} ==> x \ {a<.. (x = a) \ (x = b)" - unfolding vec_eq_iff less_vec_def less_eq_vec_def mem_interval_cart - by (auto simp del:dest_vec1_eq) - -lemma in_interval_1: - fixes x :: "real^1" - shows "(x \ {a .. b} \ dest_vec1 a \ dest_vec1 x \ dest_vec1 x \ dest_vec1 b) \ - (x \ {a<.. dest_vec1 a < dest_vec1 x \ dest_vec1 x < dest_vec1 b)" - unfolding vec_eq_iff less_vec_def less_eq_vec_def mem_interval_cart - by (auto simp del:dest_vec1_eq) - -lemma interval_eq_empty_1: - fixes a :: "real^1" - shows "{a .. b} = {} \ dest_vec1 b < dest_vec1 a" - and "{a<.. dest_vec1 b \ dest_vec1 a" - unfolding interval_eq_empty_cart and ex_1 by auto - -lemma subset_interval_1: - fixes a :: "real^1" - shows "({a .. b} \ {c .. d} \ dest_vec1 b < dest_vec1 a \ - dest_vec1 c \ dest_vec1 a \ dest_vec1 a \ dest_vec1 b \ dest_vec1 b \ dest_vec1 d)" - "({a .. b} \ {c<.. dest_vec1 b < dest_vec1 a \ - dest_vec1 c < dest_vec1 a \ dest_vec1 a \ dest_vec1 b \ dest_vec1 b < dest_vec1 d)" - "({a<.. {c .. d} \ dest_vec1 b \ dest_vec1 a \ - dest_vec1 c \ dest_vec1 a \ dest_vec1 a < dest_vec1 b \ dest_vec1 b \ dest_vec1 d)" - "({a<.. {c<.. dest_vec1 b \ dest_vec1 a \ - dest_vec1 c \ dest_vec1 a \ dest_vec1 a < dest_vec1 b \ dest_vec1 b \ dest_vec1 d)" - unfolding subset_interval_cart[of a b c d] unfolding forall_1 by auto - -lemma eq_interval_1: - fixes a :: "real^1" - shows "{a .. b} = {c .. d} \ - dest_vec1 b < dest_vec1 a \ dest_vec1 d < dest_vec1 c \ - dest_vec1 a = dest_vec1 c \ dest_vec1 b = dest_vec1 d" - unfolding set_eq_subset[of "{a .. b}" "{c .. d}"] - unfolding subset_interval_1(1)[of a b c d] - unfolding subset_interval_1(1)[of c d a b] - by auto - -lemma disjoint_interval_1: - fixes a :: "real^1" - shows - "{a .. b} \ {c .. d} = {} \ - dest_vec1 b < dest_vec1 a \ dest_vec1 d < dest_vec1 c \ dest_vec1 b < dest_vec1 c \ dest_vec1 d < dest_vec1 a" - "{a .. b} \ {c<.. - dest_vec1 b < dest_vec1 a \ dest_vec1 d \ dest_vec1 c \ dest_vec1 b \ dest_vec1 c \ dest_vec1 d \ dest_vec1 a" - "{a<.. {c .. d} = {} \ - dest_vec1 b \ dest_vec1 a \ dest_vec1 d < dest_vec1 c \ dest_vec1 b \ dest_vec1 c \ dest_vec1 d \ dest_vec1 a" - "{a<.. {c<.. - dest_vec1 b \ dest_vec1 a \ dest_vec1 d \ dest_vec1 c \ dest_vec1 b \ dest_vec1 c \ dest_vec1 d \ dest_vec1 a" - unfolding disjoint_interval_cart and ex_1 by auto - -lemma open_closed_interval_1: - fixes a :: "real^1" - shows "{a<.. dest_vec1 b ==> {a .. b} = {a<.. {a,b}" - unfolding set_eq_iff - apply simp - unfolding less_vec_def and less_eq_vec_def and forall_1 and dest_vec1_eq[symmetric] - apply (auto simp del:dest_vec1_eq) - done - -lemma Lim_drop_le: - fixes f :: "'a \ real^1" - shows "(f ---> l) net \ \ trivial_limit net \ - eventually (\x. dest_vec1 (f x) \ b) net ==> dest_vec1 l \ b" - using Lim_component_le_cart[of f l net 1 b] by auto - -lemma Lim_drop_ge: - fixes f :: "'a \ real^1" - shows "(f ---> l) net \ \ trivial_limit net \ - eventually (\x. b \ dest_vec1 (f x)) net ==> b \ dest_vec1 l" - using Lim_component_ge_cart[of f l net b 1] by auto - - -text{* Also more convenient formulations of monotone convergence. *} - -lemma bounded_increasing_convergent: - fixes s :: "nat \ real^1" - assumes "bounded {s n| n::nat. True}" "\n. dest_vec1(s n) \ dest_vec1(s(Suc n))" - shows "\l. (s ---> l) sequentially" -proof - - obtain a where a:"\n. \dest_vec1 (s n)\ \ a" - using assms(1)[unfolded bounded_iff abs_dest_vec1] by auto - { fix m::nat - have "\ n. n\m \ dest_vec1 (s m) \ dest_vec1 (s n)" - apply (induct_tac n) - apply simp - using assms(2) apply (erule_tac x="na" in allE) - apply (auto simp add: not_less_eq_eq) - done - } - hence "\m n. m \ n \ dest_vec1 (s m) \ dest_vec1 (s n)" - by auto - then obtain l where "\e>0. \N. \n\N. \dest_vec1 (s n) - l\ < e" - using convergent_bounded_monotone[OF a] unfolding monoseq_def by auto - thus ?thesis unfolding LIMSEQ_def apply(rule_tac x="vec1 l" in exI) - unfolding dist_norm unfolding abs_dest_vec1 by auto -qed - -lemma dest_vec1_simps[simp]: - fixes a :: "real^1" - shows "a$1 = 0 \ a = 0" (*"a \ 1 \ dest_vec1 a \ 1" "0 \ a \ 0 \ dest_vec1 a"*) - "a \ b \ dest_vec1 a \ dest_vec1 b" "dest_vec1 (1::real^1) = 1" - by (auto simp add: less_eq_vec_def vec_eq_iff) - -lemma dest_vec1_inverval: - "dest_vec1 ` {a .. b} = {dest_vec1 a .. dest_vec1 b}" - "dest_vec1 ` {a<.. b} = {dest_vec1 a<.. dest_vec1 b}" - "dest_vec1 ` {a ..x. dest_vec1 (f x)) S" - using dest_vec1_sum[OF assms] by auto - -lemma open_dest_vec1_vimage: "open S \ open (dest_vec1 -` S)" - unfolding open_vec_def forall_1 by auto - -lemma tendsto_dest_vec1 [tendsto_intros]: - "(f ---> l) net \ ((\x. dest_vec1 (f x)) ---> dest_vec1 l) net" - by (rule tendsto_vec_nth) - -lemma continuous_dest_vec1: - "continuous net f \ continuous net (\x. dest_vec1 (f x))" - unfolding continuous_def by (rule tendsto_dest_vec1) - -lemma forall_dest_vec1: "(\x. P x) \ (\x. P(dest_vec1 x))" - apply safe defer - apply (erule_tac x="vec1 x" in allE) - apply auto - done - -lemma forall_of_dest_vec1: "(\v. P (\x. dest_vec1 (v x))) \ (\x. P x)" - apply rule - apply rule - apply (erule_tac x="vec1 \ x" in allE) - unfolding o_def vec1_dest_vec1 - apply auto - done - -lemma forall_of_dest_vec1': "(\v. P (dest_vec1 v)) \ (\x. P x)" - apply rule - apply rule - apply (erule_tac x="(vec1 x)" in allE) defer - apply rule - apply (erule_tac x="dest_vec1 v" in allE) - unfolding o_def vec1_dest_vec1 - apply auto - done - -lemma dist_vec1_0[simp]: "dist(vec1 (x::real)) 0 = norm x" - unfolding dist_norm by auto - -lemma bounded_linear_vec1_dest_vec1: - fixes f :: "real \ real" - shows "linear (vec1 \ f \ dest_vec1) = bounded_linear f" (is "?l = ?r") -proof - - { assume ?l - then have "\K. \x. norm ((vec1 \ f \ dest_vec1) x) \ K * norm x" by (rule linear_bounded) - then guess K .. - hence "\K. \x. \f x\ \ \x\ * K" - apply(rule_tac x=K in exI) - unfolding vec1_dest_vec1_simps by (auto simp add:field_simps) - } - thus ?thesis - unfolding linear_def bounded_linear_def additive_def bounded_linear_axioms_def o_def - unfolding vec1_dest_vec1_simps by auto -qed - -lemma vec1_le[simp]: fixes a :: real shows "vec1 a \ vec1 b \ a \ b" - unfolding less_eq_vec_def by auto -lemma vec1_less[simp]: fixes a :: real shows "vec1 a < vec1 b \ a < b" - unfolding less_vec_def by auto - - -subsection {* Derivatives on real = Derivatives on @{typ "real^1"} *} - -lemma has_derivative_within_vec1_dest_vec1: - fixes f :: "real \ real" - shows "((vec1 \ f \ dest_vec1) has_derivative (vec1 \ f' \ dest_vec1)) (at (vec1 x) within vec1 ` s) - = (f has_derivative f') (at x within s)" - unfolding has_derivative_within - unfolding bounded_linear_vec1_dest_vec1[unfolded linear_conv_bounded_linear] - unfolding o_def Lim_within Ball_def unfolding forall_vec1 - unfolding vec1_dest_vec1_simps dist_vec1_0 image_iff - by auto - -lemma has_derivative_at_vec1_dest_vec1: - fixes f :: "real \ real" - shows "((vec1 \ f \ dest_vec1) has_derivative (vec1 \ f' \ dest_vec1)) (at (vec1 x)) = (f has_derivative f') (at x)" - using has_derivative_within_vec1_dest_vec1[where s=UNIV, unfolded range_vec1 within_UNIV] - by auto - -lemma bounded_linear_vec1': - fixes f :: "'a::real_normed_vector\real" - shows "bounded_linear f = bounded_linear (vec1 \ f)" - unfolding bounded_linear_def additive_def bounded_linear_axioms_def o_def - unfolding vec1_dest_vec1_simps by auto - -lemma bounded_linear_dest_vec1: - fixes f :: "real\'a::real_normed_vector" - shows "bounded_linear f = bounded_linear (f \ dest_vec1)" - unfolding bounded_linear_def additive_def bounded_linear_axioms_def o_def - unfolding vec1_dest_vec1_simps - by auto - -lemma has_derivative_at_vec1: - fixes f :: "'a::real_normed_vector\real" - shows "(f has_derivative f') (at x) = ((vec1 \ f) has_derivative (vec1 \ f')) (at x)" - unfolding has_derivative_at - unfolding bounded_linear_vec1'[unfolded linear_conv_bounded_linear] - unfolding o_def Lim_at - unfolding vec1_dest_vec1_simps dist_vec1_0 - by auto - -lemma has_derivative_within_dest_vec1: - fixes f :: "real\'a::real_normed_vector" - shows "((f \ dest_vec1) has_derivative (f' \ dest_vec1)) (at (vec1 x) within vec1 ` s) = - (f has_derivative f') (at x within s)" - unfolding has_derivative_within bounded_linear_dest_vec1 - unfolding o_def Lim_within Ball_def - unfolding vec1_dest_vec1_simps dist_vec1_0 image_iff - by auto - -lemma has_derivative_at_dest_vec1: - fixes f :: "real\'a::real_normed_vector" - shows "((f \ dest_vec1) has_derivative (f' \ dest_vec1)) (at (vec1 x)) = - (f has_derivative f') (at x)" - using has_derivative_within_dest_vec1[where s=UNIV] by simp - - -subsection {* In particular if we have a mapping into @{typ "real^1"}. *} - -lemma onorm_vec1: - fixes f::"real \ real" - shows "onorm (\x. vec1 (f (dest_vec1 x))) = onorm f" -proof - - have "\x::real^1. norm x = 1 \ x\{vec1 -1, vec1 (1::real)}" - unfolding forall_vec1 by (auto simp add: vec_eq_iff) - hence 1: "{x. norm x = 1} = {vec1 -1, vec1 (1::real)}" by auto - have 2: "{norm (vec1 (f (dest_vec1 x))) |x. norm x = 1} = - (\x. norm (vec1 (f (dest_vec1 x)))) ` {x. norm x=1}" - by auto - have "\x::real. norm x = 1 \ x\{-1, 1}" by auto - hence 3:"{x. norm x = 1} = {-1, (1::real)}" by auto - have 4:"{norm (f x) |x. norm x = 1} = (\x. norm (f x)) ` {x. norm x=1}" by auto - show ?thesis - unfolding onorm_def 1 2 3 4 by (simp add:Sup_finite_Max) -qed - -lemma convex_vec1:"convex (vec1 ` s) = convex (s::real set)" - unfolding convex_def Ball_def forall_vec1 - unfolding vec1_dest_vec1_simps image_iff - by auto - lemma bounded_linear_component_cart[intro]: "bounded_linear (\x::real^'n. x $ k)" apply (rule bounded_linearI[where K=1]) using component_le_norm_cart[of _ k] unfolding real_norm_def by auto -lemma bounded_vec1[intro]: "bounded s \ bounded (vec1 ` (s::real set))" - unfolding bounded_def apply safe apply(rule_tac x="vec1 x" in exI,rule_tac x=e in exI) - apply (auto simp add: dist_real dist_real_def) - done - -(*lemma content_closed_interval_cases_cart: - "content {a..b::real^'n} = - (if {a..b} = {} then 0 else setprod (\i. b$i - a$i) UNIV)" -proof(cases "{a..b} = {}") - case True thus ?thesis unfolding content_def by auto -next case Falsethus ?thesis unfolding content_def unfolding if_not_P[OF False] - proof(cases "\i. a $ i \ b $ i") - case False thus ?thesis unfolding content_def using t by auto - next case True note interval_eq_empty - apply auto - - sorry*) - lemma integral_component_eq_cart[simp]: fixes f :: "'n::ordered_euclidean_space \ real^'m" assumes "f integrable_on s" @@ -2309,39 +1348,8 @@ unfolding vec_lambda_beta by auto -(*lemma content_split_cart: - "content {a..b::real^'n} = content({a..b} \ {x. x$k \ c}) + content({a..b} \ {x. x$k >= c})" -proof- note simps = interval_split_cart content_closed_interval_cases_cart vec_lambda_beta less_eq_vec_def - { presume "a\b \ ?thesis" thus ?thesis apply(cases "a\b") unfolding simps by auto } - have *:"UNIV = insert k (UNIV - {k})" "\x. finite (UNIV-{x::'n})" "\x. x\UNIV-{x}" by auto - have *:"\X Y Z. (\i\UNIV. Z i (if i = k then X else Y i)) = Z k X * (\i\UNIV-{k}. Z i (Y i))" - "(\i\UNIV. b$i - a$i) = (\i\UNIV-{k}. b$i - a$i) * (b$k - a$k)" - apply(subst *(1)) defer apply(subst *(1)) unfolding setprod_insert[OF *(2-)] by auto - assume as:"a\b" moreover have "\x. min (b $ k) c = max (a $ k) c - \ x* (b$k - a$k) = x*(max (a $ k) c - a $ k) + x*(b $ k - max (a $ k) c)" - by (auto simp add:field_simps) - moreover have "\ a $ k \ c \ \ c \ b $ k \ False" - unfolding not_le using as[unfolded less_eq_vec_def,rule_format,of k] by auto - ultimately show ?thesis - unfolding simps unfolding *(1)[of "\i x. b$i - x"] *(1)[of "\i x. x - a$i"] *(2) by(auto) -qed*) - -lemma has_integral_vec1: - assumes "(f has_integral k) {a..b}" - shows "((\x. vec1 (f x)) has_integral (vec1 k)) {a..b}" -proof - - have *: "\p. (\(x, k)\p. content k *\<^sub>R vec1 (f x)) - vec1 k = - vec1 ((\(x, k)\p. content k *\<^sub>R f x) - k)" - unfolding vec_sub vec_eq_iff by (auto simp add: split_beta) - show ?thesis - using assms unfolding has_integral - apply safe - apply(erule_tac x=e in allE,safe) - apply(rule_tac x=d in exI,safe) - apply(erule_tac x=p in allE,safe) - unfolding * norm_vector_1 - apply auto - done -qed +lemma interval_bij_bij_cart: fixes x::"real^'n" assumes "\i. a$i < b$i \ u$i < v$i" + shows "interval_bij (a,b) (u,v) (interval_bij (u,v) (a,b) x) = x" + using assms by (intro interval_bij_bij) (auto simp: Basis_vec_def inner_axis) end diff -r 46be26e02456 -r 899c9c4e4a4c src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy --- a/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy Fri Dec 14 14:46:01 2012 +0100 +++ b/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy Fri Dec 14 15:46:01 2012 +0100 @@ -101,56 +101,22 @@ lemma span_eq[simp]: "(span s = s) <-> subspace s" unfolding span_def by (rule hull_eq, rule subspace_Inter) -lemma basis_inj_on: "d \ {.. inj_on (basis :: nat => 'n::euclidean_space) d" - by (auto simp add: inj_on_def euclidean_eq[where 'a='n]) - -lemma finite_substdbasis: "finite {basis i ::'n::euclidean_space |i. i : (d:: nat set)}" (is "finite ?S") -proof - - have eq: "?S = basis ` d" by blast - show ?thesis - unfolding eq - apply (rule finite_subset[OF _ range_basis_finite]) - apply auto - done -qed - -lemma card_substdbasis: - assumes "d \ {..{..R basis i) d = (x::'a::euclidean_space) - <-> (!i f i = x$$i) & (i ~: d --> x $$ i = 0))" + assumes d: "d \ Basis" + shows "(\i\d. f i *\<^sub>R i) = (x::'a::euclidean_space) + \ (\i\Basis. (i \ d \ f i = x \ i) \ (i \ d \ x \ i = 0))" proof - have *: "\x a b P. x * (if P then a else b) = (if P then x*a else x*b)" by auto - have **: "finite d" apply(rule finite_subset[OF assms]) by fastforce - have ***: "\i. (setsum (%i. f i *\<^sub>R ((basis i)::'a)) d) $$ i = (\x\d. if x = i then f x else 0)" - unfolding euclidean_component_setsum euclidean_component_scaleR basis_component * - apply (rule setsum_cong2) - using assms apply auto - done + have **: "finite d" by (auto intro: finite_subset[OF assms]) + have ***: "\i. i \ Basis \ (\i\d. f i *\<^sub>R i) \ i = (\x\d. if x = i then f x else 0)" + using d + by (auto intro!: setsum_cong simp: inner_Basis inner_setsum_left) show ?thesis - unfolding euclidean_eq[where 'a='a] *** setsum_delta[OF **] using assms by auto -qed - -lemma independent_substdbasis: - assumes "d \ {..i. i convex {x. P i x}" - shows "convex {x. \ii x$$i)}" + assumes "\i. i\Basis \ convex {x. P i x}" + shows "convex {x. \i\Basis. P i (x\i)}" + using assms unfolding convex_def + by (auto simp: inner_add_left) + +lemma convex_positive_orthant: "convex {x::'a::euclidean_space. (\i\Basis. 0 \ x\i)}" by (rule convex_box) (simp add: atLeast_def[symmetric] convex_real_interval) lemma convex_local_global_minimum: @@ -2073,40 +2040,39 @@ from this show ?thesis using assms `span B=S` by auto qed -lemma span_substd_basis: assumes "d\{.. x$$i = 0)}" - (is "span ?A = ?B") +lemma span_substd_basis: + assumes d: "d \ Basis" + shows "span d = {x. \i\Basis. i \ d \ x\i = 0}" (is "_ = ?B") proof- -have "?A <= ?B" by auto +have "d <= ?B" using d by (auto simp: inner_Basis) moreover have s: "subspace ?B" using subspace_substandard[of "%i. i ~: d"] . -ultimately have "span ?A <= ?B" using span_mono[of "?A" "?B"] span_eq[of "?B"] by blast -moreover have "card d <= dim (span ?A)" using independent_card_le_dim[of "?A" "span ?A"] - independent_substdbasis[OF assms] card_substdbasis[OF assms] span_inc[of "?A"] by auto -moreover hence "dim ?B <= dim (span ?A)" using dim_substandard[OF assms] by auto -ultimately show ?thesis using s subspace_dim_equal[of "span ?A" "?B"] - subspace_span[of "?A"] by auto +ultimately have "span d <= ?B" using span_mono[of d "?B"] span_eq[of "?B"] by blast +moreover have "card d <= dim (span d)" using independent_card_le_dim[of d "span d"] + independent_substdbasis[OF assms] span_inc[of d] by auto +moreover hence "dim ?B <= dim (span d)" using dim_substandard[OF assms] by auto +ultimately show ?thesis using s subspace_dim_equal[of "span d" "?B"] + subspace_span[of d] by auto qed lemma basis_to_substdbasis_subspace_isomorphism: fixes B :: "('a::euclidean_space) set" assumes "independent B" -shows "EX f d. card d = card B & linear f & f ` B = {basis i::'a |i. i : (d :: nat set)} & - f ` span B = {x. ALL i x $$ i = (0::real)} & inj_on f (span B) \ d\{.. linear f \ f ` B = d \ + f ` span B = {x. \i\Basis. i \ d \ x \ i = 0} \ inj_on f (span B) \ d \ Basis" proof- have B:"card B=dim B" using dim_unique[of B B "card B"] assms span_inc[of B] by auto - def d \ "{..{.. card (Basis :: 'a set)" using dim_subset_UNIV[of B] by simp + from ex_card[OF this] obtain d :: "'a set" where d: "d \ Basis" and t: "card d = dim B" by auto + let ?t = "{x::'a::euclidean_space. \i\Basis. i ~: d --> x\i = 0}" + have "EX f. linear f & f ` B = d & f ` span B = ?t & inj_on f (span B)" + apply (rule basis_to_basis_subspace_isomorphism[of "span B" ?t B "d"]) apply(rule subspace_span) apply(rule subspace_substandard) defer apply(rule span_inc) apply(rule assms) defer unfolding dim_span[of B] apply(rule B) - unfolding span_substd_basis[OF d,symmetric] card_substdbasis[OF d] apply(rule span_inc) + unfolding span_substd_basis[OF d, symmetric] + apply(rule span_inc) apply(rule independent_substdbasis[OF d]) apply(rule,assumption) unfolding t[symmetric] span_substd_basis[OF d] dim_substandard[OF d] by auto - from this t `card B=dim B` show ?thesis using d by auto + with t `card B = dim B` d show ?thesis by auto qed lemma aff_dim_empty: @@ -2492,7 +2458,7 @@ using as affine_affine_hull[of S] mem_affine[of "affine hull S" y x "(1 / e)" "-((1 - e) / e)"] by (simp add: algebra_simps) have "dist c ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) = abs(1/e) * norm (e *\<^sub>R c - y + (1 - e) *\<^sub>R x)" unfolding dist_norm unfolding norm_scaleR[symmetric] apply(rule arg_cong[where f=norm]) using `e>0` - by(auto simp add:euclidean_eq[where 'a='a] field_simps) + by(auto simp add:euclidean_eq_iff[where 'a='a] field_simps inner_simps) also have "... = abs(1/e) * norm (x - e *\<^sub>R (x - c) - y)" by(auto intro!:arg_cong[where f=norm] simp add: algebra_simps) also have "... < d" using as[unfolded dist_norm] and `e>0` by(auto simp add:pos_divide_less_eq[OF `e>0`] mult_commute) @@ -2770,9 +2736,9 @@ subsection{* Some Properties of subset of standard basis *} -lemma affine_hull_substd_basis: assumes "d\{.. x$$i = 0)}" +lemma affine_hull_substd_basis: assumes "d\Basis" + shows "affine hull (insert 0 d) = + {x::'a::euclidean_space. (\i\Basis. i ~: d --> x\i = 0)}" (is "affine hull (insert 0 ?A) = ?B") proof- have *:"\A. op + (0\'a) ` A = A" "\A. op + (- (0\'a)) ` A = A" by auto show ?thesis unfolding affine_hull_insert_span_gen span_substd_basis[OF assms,symmetric] * .. @@ -3230,10 +3196,10 @@ assumes "convex (s::('a::euclidean_space) set)" "closed s" "0 \ s" shows "\a b. a \ 0 \ 0 < b \ (\x\s. inner a x > b)" proof(cases "s={}") - case True have "norm ((basis 0)::'a) = 1" by auto - hence "norm ((basis 0)::'a) = 1" "basis 0 \ (0::'a)" defer - apply(subst norm_le_zero_iff[symmetric]) by auto - thus ?thesis apply(rule_tac x="basis 0" in exI, rule_tac x=1 in exI) + case True + have "norm ((SOME i. i\Basis)::'a) = 1" "(SOME i. i\Basis) \ (0::'a)" defer + apply(subst norm_le_zero_iff[symmetric]) by (auto simp: SOME_Basis) + thus ?thesis apply(rule_tac x="SOME i. i\Basis" in exI, rule_tac x=1 in exI) using True using DIM_positive[where 'a='a] by auto next case False thus ?thesis using False using separating_hyperplane_closed_point[OF assms] apply - apply(erule exE)+ unfolding inner_zero_right apply(rule_tac x=a in exI, rule_tac x=b in exI) by auto qed @@ -3703,10 +3669,10 @@ case False thus ?thesis apply (intro continuous_intros) using cont_surfpi unfolding continuous_on_eq_continuous_at[OF open_delete[OF open_UNIV]] o_def by auto next obtain B where B:"\x\s. norm x \ B" using compact_imp_bounded[OF assms(1)] unfolding bounded_iff by auto - hence "B > 0" using assms(2) unfolding subset_eq apply(erule_tac x="basis 0" in ballE) defer - apply(erule_tac x="basis 0" in ballE) + hence "B > 0" using assms(2) unfolding subset_eq apply(erule_tac x="SOME i. i\Basis" in ballE) defer + apply(erule_tac x="SOME i. i\Basis" in ballE) unfolding Ball_def mem_cball dist_norm using DIM_positive[where 'a='a] - by auto + by (auto simp: SOME_Basis) case True show ?thesis unfolding True continuous_at Lim_at apply(rule,rule) apply(rule_tac x="e / B" in exI) apply(rule) apply(rule divide_pos_pos) prefer 3 apply(rule,rule,erule conjE) unfolding norm_zero scaleR_zero_left dist_norm diff_0_right norm_scaleR abs_norm_cancel proof- @@ -3849,7 +3815,8 @@ hence "a < b" unfolding * using as(4) apply(rule_tac mult_left_less_imp_less) by(auto simp add: field_simps) hence "u * a + v * b \ b" unfolding ** using **(2) as(3) by(auto simp add: field_simps intro!:mult_right_mono) } ultimately show "u *\<^sub>R x + v *\<^sub>R y \ s" apply- apply(rule assms[unfolded is_interval_def, rule_format, OF as(1,2)]) - using as(3-) DIM_positive[where 'a='a] by auto qed + using as(3-) DIM_positive[where 'a='a] by (auto simp: inner_simps) +qed lemma is_interval_connected: fixes s :: "('a::euclidean_space) set" @@ -3892,8 +3859,8 @@ subsection {* Another intermediate value theorem formulation *} lemma ivt_increasing_component_on_1: fixes f::"real \ 'a::euclidean_space" - assumes "a \ b" "continuous_on {a .. b} f" "(f a)$$k \ y" "y \ (f b)$$k" - shows "\x\{a..b}. (f x)$$k = y" + assumes "a \ b" "continuous_on {a .. b} f" "(f a)\k \ y" "y \ (f b)\k" + shows "\x\{a..b}. (f x)\k = y" proof- have "f a \ f ` {a..b}" "f b \ f ` {a..b}" apply(rule_tac[!] imageI) using assms(1) by auto thus ?thesis using connected_ivt_component[of "f ` {a..b}" "f a" "f b" k y] @@ -3902,20 +3869,20 @@ lemma ivt_increasing_component_1: fixes f::"real \ 'a::euclidean_space" shows "a \ b \ \x\{a .. b}. continuous (at x) f - \ f a$$k \ y \ y \ f b$$k \ \x\{a..b}. (f x)$$k = y" + \ f a\k \ y \ y \ f b\k \ \x\{a..b}. (f x)\k = y" by(rule ivt_increasing_component_on_1) (auto simp add: continuous_at_imp_continuous_on) lemma ivt_decreasing_component_on_1: fixes f::"real \ 'a::euclidean_space" - assumes "a \ b" "continuous_on {a .. b} f" "(f b)$$k \ y" "y \ (f a)$$k" - shows "\x\{a..b}. (f x)$$k = y" + assumes "a \ b" "continuous_on {a .. b} f" "(f b)\k \ y" "y \ (f a)\k" + shows "\x\{a..b}. (f x)\k = y" apply(subst neg_equal_iff_equal[symmetric]) using ivt_increasing_component_on_1[of a b "\x. - f x" k "- y"] using assms using continuous_on_minus by auto lemma ivt_decreasing_component_1: fixes f::"real \ 'a::euclidean_space" shows "a \ b \ \x\{a .. b}. continuous (at x) f - \ f b$$k \ y \ y \ f a$$k \ \x\{a..b}. (f x)$$k = y" + \ f b\k \ y \ y \ f a\k \ \x\{a..b}. (f x)\k = y" by(rule ivt_decreasing_component_on_1) (auto simp: continuous_at_imp_continuous_on) @@ -3933,104 +3900,127 @@ thus "f x \ b" using assms(1)[unfolded convex_on[OF convex_convex_hull], rule_format, of k u v] unfolding obt(2-3) using obt(1) and hull_subset[unfolded subset_eq, rule_format, of _ s] by auto qed +lemma inner_setsum_Basis[simp]: "\i. i \ Basis \ (\Basis) \ i = 1" + by (simp add: One_def inner_setsum_left setsum_cases inner_Basis) + lemma unit_interval_convex_hull: - "{0::'a::ordered_euclidean_space .. (\\ i. 1)} = convex hull {x. \i (x$$i = 1)}" + defines "One \ (\Basis)" + shows "{0::'a::ordered_euclidean_space .. One} = + convex hull {x. \i\Basis. (x\i = 0) \ (x\i = 1)}" (is "?int = convex hull ?points") -proof- have 01:"{0,(\\ i. 1)} \ convex hull ?points" apply rule apply(rule_tac hull_subset[unfolded subset_eq, rule_format]) by auto - { fix n x assume "x\{0::'a::ordered_euclidean_space .. \\ i. 1}" "n \ DIM('a)" "card {i. i x$$i \ 0} \ n" +proof - + have One[simp]: "\i. i \ Basis \ One \ i = 1" + by (simp add: One_def inner_setsum_left setsum_cases inner_Basis) + have 01:"{0,One} \ convex hull ?points" + apply rule apply(rule_tac hull_subset[unfolded subset_eq, rule_format]) by auto + { fix n x assume "x\{0::'a::ordered_euclidean_space .. One}" "n \ DIM('a)" "card {i. i\Basis \ x\i \ 0} \ n" hence "x\convex hull ?points" proof(induct n arbitrary: x) - case 0 hence "x = 0" apply(subst euclidean_eq) apply rule by auto + case 0 hence "x = 0" apply(subst euclidean_eq_iff) apply rule by auto thus "x\convex hull ?points" using 01 by auto next - case (Suc n) show "x\convex hull ?points" proof(cases "{i. i x$$i \ 0} = {}") - case True hence "x = 0" apply(subst euclidean_eq) by auto + case (Suc n) show "x\convex hull ?points" proof(cases "{i. i\Basis \ x\i \ 0} = {}") + case True hence "x = 0" apply(subst euclidean_eq_iff) by auto thus "x\convex hull ?points" using 01 by auto next - case False def xi \ "Min ((\i. x$$i) ` {i. i x$$i \ 0})" - have "xi \ (\i. x$$i) ` {i. i x$$i \ 0}" unfolding xi_def apply(rule Min_in) using False by auto - then obtain i where i':"x$$i = xi" "x$$i \ 0" "ij. j x$$j > 0 \ x$$i \ x$$j" + case False def xi \ "Min ((\i. x\i) ` {i. i\Basis \ x\i \ 0})" + have "xi \ (\i. x\i) ` {i. i\Basis \ x\i \ 0}" unfolding xi_def apply(rule Min_in) using False by auto + then obtain i where i':"x\i = xi" "x\i \ 0" "i\Basis" by auto + have i:"\j. j\Basis \ x\j > 0 \ x\i \ x\j" unfolding i'(1) xi_def apply(rule_tac Min_le) unfolding image_iff defer apply(rule_tac x=j in bexI) using i' by auto - have i01:"x$$i \ 1" "x$$i > 0" using Suc(2)[unfolded mem_interval,rule_format,of i] - using i'(2-) `x$$i \ 0` by auto - show ?thesis proof(cases "x$$i=1") - case True have "\j\{i. i x$$i \ 0}. x$$j = 1" apply(rule, rule ccontr) unfolding mem_Collect_eq - proof(erule conjE) fix j assume as:"x $$ j \ 0" "x $$ j \ 1" "j {0<..<1}" using Suc(2) by(auto simp add: eucl_le[where 'a='a] elim!:allE[where x=j]) - hence "x$$j \ op $$ x ` {i. i x $$ i \ 0}" using as(3) by auto - hence "x$$j \ x$$i" unfolding i'(1) xi_def apply(rule_tac Min_le) by auto + have i01:"x\i \ 1" "x\i > 0" using Suc(2)[unfolded mem_interval,rule_format,of i] + using i'(2-) `x\i \ 0` by auto + show ?thesis proof(cases "x\i=1") + case True have "\j\{i. i\Basis \ x\i \ 0}. x\j = 1" apply(rule, rule ccontr) unfolding mem_Collect_eq + proof(erule conjE) fix j assume as:"x \ j \ 0" "x \ j \ 1" "j\Basis" + hence j:"x\j \ {0<..<1}" using Suc(2) + by (auto simp add: eucl_le[where 'a='a] elim!:allE[where x=j]) + hence "x\j \ op \ x ` {i. i\Basis \ x \ i \ 0}" using as(3) by auto + hence "x\j \ x\i" unfolding i'(1) xi_def apply(rule_tac Min_le) by auto thus False using True Suc(2) j by(auto simp add: elim!:ballE[where x=j]) qed thus "x\convex hull ?points" apply(rule_tac hull_subset[unfolded subset_eq, rule_format]) by auto - next let ?y = "\j. if x$$j = 0 then 0 else (x$$j - x$$i) / (1 - x$$i)" - case False hence *:"x = x$$i *\<^sub>R (\\ j. if x$$j = 0 then 0 else 1) + (1 - x$$i) *\<^sub>R (\\ j. ?y j)" - apply(subst euclidean_eq) by(auto simp add: field_simps) - { fix j assume j:"j 0 \ 0 \ (x $$ j - x $$ i) / (1 - x $$ i)" "(x $$ j - x $$ i) / (1 - x $$ i) \ 1" + next + let ?y = "\j\Basis. (if x\j = 0 then 0 else (x\j - x\i) / (1 - x\i)) *\<^sub>R j" + case False + then have *: "x = (x\i) *\<^sub>R (\j\Basis. (if x\j = 0 then 0 else 1) *\<^sub>R j) + (1 - x\i) *\<^sub>R ?y" + by (subst euclidean_eq_iff) (simp add: inner_simps) + { fix j :: 'a assume j:"j\Basis" + have "x\j \ 0 \ 0 \ (x \ j - x \ i) / (1 - x \ i)" "(x \ j - x \ i) / (1 - x \ i) \ 1" apply(rule_tac divide_nonneg_pos) using i(1)[of j] using False i01 using Suc(2)[unfolded mem_interval, rule_format, of j] using j - by(auto simp add:field_simps) - hence "0 \ ?y j \ ?y j \ 1" by auto } - moreover have "i\{j. j x$$j \ 0} - {j. j ((\\ j. ?y j)::'a) $$ j \ 0}" + by(auto simp add: field_simps) + with j have "0 \ ?y \ j \ ?y \ j \ 1" by (auto simp: inner_simps) } + moreover have "i\{j. j\Basis \ x\j \ 0} - {j. j\Basis \ ?y \ j \ 0}" using i01 using i'(3) by auto - hence "{j. j x$$j \ 0} \ {j. j ((\\ j. ?y j)::'a) $$ j \ 0}" using i'(3) by blast - hence **:"{j. j ((\\ j. ?y j)::'a) $$ j \ 0} \ {j. j x$$j \ 0}" apply - apply rule + hence "{j. j\Basis \ x\j \ 0} \ {j. j\Basis \ ?y \ j \ 0}" using i'(3) by blast + hence **:"{j. j\Basis \ ?y \ j \ 0} \ {j. j\Basis \ x\j \ 0}" by auto - have "card {j. j ((\\ j. ?y j)::'a) $$ j \ 0} \ n" + have "card {j. j\Basis \ ?y \ j \ 0} \ n" using less_le_trans[OF psubset_card_mono[OF _ **] Suc(4)] by auto - ultimately show ?thesis apply(subst *) apply(rule convex_convex_hull[unfolded convex_def, rule_format]) - apply(rule_tac hull_subset[unfolded subset_eq, rule_format]) defer apply(rule Suc(1)) - unfolding mem_interval using i01 Suc(3) by auto - qed qed qed } note * = this - have **:"DIM('a) = card {..\ i. 1)::'a::ordered_euclidean_space} = convex hull s" - apply(rule that[of "{x::'a. \i x$$i=1}"]) - apply(rule finite_subset[of _ "(\s. (\\ i. if i\s then 1::real else 0)::'a) ` Pow {..Basis} = convex hull s" + apply(rule that[of "{x::'a. \i\Basis. x\i=0 \ x\i=1}"]) + apply(rule finite_subset[of _ "(\s. (\i\Basis. (if i\s then 1 else 0) *\<^sub>R i)::'a) ` Pow Basis"]) prefer 3 apply(rule unit_interval_convex_hull) apply rule unfolding mem_Collect_eq proof- - fix x::"'a" assume as:"\i x $$ i = 1" - show "x \ (\s. \\ i. if i \ s then 1 else 0) ` Pow {..i\Basis. x \ i = 0 \ x \ i = 1" + show "x \ (\s. \i\Basis. (if i\s then 1 else 0) *\<^sub>R i) ` Pow Basis" + apply(rule image_eqI[where x="{i. i\Basis \ x\i = 1}"]) + using as apply(subst euclidean_eq_iff) by (auto simp: inner_setsum_left_Basis) +qed auto text {* Hence any cube (could do any nonempty interval). *} lemma cube_convex_hull: assumes "0 < d" obtains s::"('a::ordered_euclidean_space) set" where - "finite s" "{x - (\\ i. d) .. x + (\\ i. d)} = convex hull s" proof- - let ?d = "(\\ i. d)::'a" - have *:"{x - ?d .. x + ?d} = (\y. x - ?d + (2 * d) *\<^sub>R y) ` {0 .. \\ i. 1}" apply(rule set_eqI, rule) + "finite s" "{x - (\i\Basis. d*\<^sub>Ri) .. x + (\i\Basis. d*\<^sub>Ri)} = convex hull s" proof- + let ?d = "(\i\Basis. d*\<^sub>Ri)::'a" + have *:"{x - ?d .. x + ?d} = (\y. x - ?d + (2 * d) *\<^sub>R y) ` {0 .. \Basis}" apply(rule set_eqI, rule) unfolding image_iff defer apply(erule bexE) proof- fix y assume as:"y\{x - ?d .. x + ?d}" - { fix i assume i:"i d + y $$ i" "y $$ i \ d + x $$ i" - using as[unfolded mem_interval, THEN spec[where x=i]] i - by auto - hence "1 \ inverse d * (x $$ i - y $$ i)" "1 \ inverse d * (y $$ i - x $$ i)" + { fix i :: 'a assume i:"i\Basis" have "x \ i \ d + y \ i" "y \ i \ d + x \ i" + using as[unfolded mem_interval, THEN bspec[where x=i]] i + by (auto simp: inner_simps) + hence "1 \ inverse d * (x \ i - y \ i)" "1 \ inverse d * (y \ i - x \ i)" apply(rule_tac[!] mult_left_le_imp_le[OF _ assms]) unfolding mult_assoc[symmetric] using assms by(auto simp add: field_simps) - hence "inverse d * (x $$ i * 2) \ 2 + inverse d * (y $$ i * 2)" - "inverse d * (y $$ i * 2) \ 2 + inverse d * (x $$ i * 2)" by(auto simp add:field_simps) } - hence "inverse (2 * d) *\<^sub>R (y - (x - ?d)) \ {0..\\ i.1}" unfolding mem_interval using assms - by(auto simp add: field_simps) - thus "\z\{0..\\ i.1}. y = x - ?d + (2 * d) *\<^sub>R z" apply- apply(rule_tac x="inverse (2 * d) *\<^sub>R (y - (x - ?d))" in bexI) + hence "inverse d * (x \ i * 2) \ 2 + inverse d * (y \ i * 2)" + "inverse d * (y \ i * 2) \ 2 + inverse d * (x \ i * 2)" by(auto simp add:field_simps) } + hence "inverse (2 * d) *\<^sub>R (y - (x - ?d)) \ {0..\Basis}" unfolding mem_interval using assms + by(auto simp add: field_simps inner_simps) + thus "\z\{0..\Basis}. y = x - ?d + (2 * d) *\<^sub>R z" apply- apply(rule_tac x="inverse (2 * d) *\<^sub>R (y - (x - ?d))" in bexI) using assms by auto next - fix y z assume as:"z\{0..\\ i.1}" "y = x - ?d + (2*d) *\<^sub>R z" - have "\i. i 0 \ d * z $$ i \ d * z $$ i \ d" - using assms as(1)[unfolded mem_interval] apply(erule_tac x=i in allE) + fix y z assume as:"z\{0..\Basis}" "y = x - ?d + (2*d) *\<^sub>R z" + have "\i. i\Basis \ 0 \ d * (z \ i) \ d * (z \ i) \ d" + using assms as(1)[unfolded mem_interval] apply(erule_tac x=i in ballE) apply rule apply(rule mult_nonneg_nonneg) prefer 3 apply(rule mult_right_le_one_le) using assms by auto thus "y \ {x - ?d..x + ?d}" unfolding as(2) mem_interval apply- apply rule using as(1)[unfolded mem_interval] - apply(erule_tac x=i in allE) using assms by auto qed - obtain s where "finite s" "{0::'a..\\ i.1} = convex hull s" using unit_cube_convex_hull by auto + apply(erule_tac x=i in ballE) using assms by (auto simp: inner_simps) qed + obtain s where "finite s" "{0::'a..\Basis} = convex hull s" using unit_cube_convex_hull by auto thus ?thesis apply(rule_tac that[of "(\y. x - ?d + (2 * d) *\<^sub>R y)` s"]) unfolding * and convex_hull_affinity by auto qed subsection {* Bounded convex function on open set is continuous *} @@ -4103,6 +4093,9 @@ subsubsection {* Hence a convex function on an open set is continuous *} +lemma real_of_nat_ge_one_iff: "1 \ real (n::nat) \ 1 \ n" + by auto + lemma convex_on_continuous: assumes "open (s::('a::ordered_euclidean_space) set)" "convex_on s f" (* FIXME: generalize to euclidean_space *) @@ -4113,29 +4106,40 @@ then obtain e where e:"cball x e \ s" "e>0" using assms(1) unfolding open_contains_cball by auto def d \ "e / real DIM('a)" have "0 < d" unfolding d_def using `e>0` dimge1 by(rule_tac divide_pos_pos, auto) - let ?d = "(\\ i. d)::'a" + let ?d = "(\i\Basis. d *\<^sub>R i)::'a" obtain c where c:"finite c" "{x - ?d..x + ?d} = convex hull c" using cube_convex_hull[OF `d>0`, of x] by auto - have "x\{x - ?d..x + ?d}" using `d>0` unfolding mem_interval by auto + have "x\{x - ?d..x + ?d}" using `d>0` unfolding mem_interval by (auto simp: inner_setsum_left_Basis inner_simps) hence "c\{}" using c by auto def k \ "Max (f ` c)" - have "convex_on {x - ?d..x + ?d} f" apply(rule convex_on_subset[OF assms(2)]) - apply(rule subset_trans[OF _ e(1)]) unfolding subset_eq mem_cball proof + have "convex_on {x - ?d..x + ?d} f" + apply(rule convex_on_subset[OF assms(2)]) + apply(rule subset_trans[OF _ e(1)]) + unfolding subset_eq mem_cball + proof fix z assume z:"z\{x - ?d..x + ?d}" - have e:"e = setsum (\i. d) {..i::'a. d) Basis" unfolding setsum_constant d_def using dimge1 unfolding real_eq_of_nat by auto show "dist x z \ e" unfolding dist_norm e apply(rule_tac order_trans[OF norm_le_l1], rule setsum_mono) - using z[unfolded mem_interval] apply(erule_tac x=i in allE) by auto qed + using z[unfolded mem_interval] apply(erule_tac x=b in ballE) by (auto simp: inner_simps) + qed hence k:"\y\{x - ?d..x + ?d}. f y \ k" unfolding c(2) apply(rule_tac convex_on_convex_hull_bound) apply assumption unfolding k_def apply(rule, rule Max_ge) using c(1) by auto - have "d \ e" unfolding d_def apply(rule mult_imp_div_pos_le) using `e>0` dimge1 unfolding mult_le_cancel_left1 by auto + have "d \ e" + unfolding d_def + apply(rule mult_imp_div_pos_le) + using `e>0` + unfolding mult_le_cancel_left1 + apply (auto simp: real_of_nat_ge_one_iff Suc_le_eq DIM_positive) + done hence dsube:"cball x d \ cball x e" unfolding subset_eq Ball_def mem_cball by auto have conv:"convex_on (cball x d) f" apply(rule convex_on_subset, rule convex_on_subset[OF assms(2)]) apply(rule e(1)) using dsube by auto hence "\y\cball x d. abs (f y) \ k + 2 * abs (f x)" apply(rule_tac convex_bounds_lemma) apply assumption proof fix y assume y:"y\cball x d" - { fix i assume "i y $$ i" "y $$ i \ x $$ i + d" - using order_trans[OF component_le_norm y[unfolded mem_cball dist_norm], of i] by auto } + { fix i :: 'a assume "i\Basis" hence "x \ i - d \ y \ i" "y \ i \ x \ i + d" + using order_trans[OF Basis_le_norm y[unfolded mem_cball dist_norm], of i] by (auto simp: inner_diff_left) } thus "f y \ k" apply(rule_tac k[rule_format]) unfolding mem_cball mem_interval dist_norm - by auto qed + by (auto simp: inner_simps) + qed hence "continuous_on (ball x d) f" apply(rule_tac convex_on_bounded_continuous) apply(rule open_ball, rule convex_on_subset[OF conv], rule ball_subset_cball) apply force @@ -4266,25 +4270,26 @@ have "norm (a - x) / norm (a - b) \ 1" unfolding divide_le_eq_1_pos[OF Fal2] unfolding as[unfolded dist_norm] norm_ge_zero by auto thus "\u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \ 0 \ u \ u \ 1" apply(rule_tac x="dist a x / dist a b" in exI) - unfolding dist_norm apply(subst euclidean_eq) apply rule defer apply(rule, rule divide_nonneg_pos) prefer 4 - proof(rule,rule) fix i assume i:"iR a + (norm (a - x) / norm (a - b)) *\<^sub>R b) $$ i = - ((norm (a - b) - norm (a - x)) * (a $$ i) + norm (a - x) * (b $$ i)) / norm (a - b)" - using Fal by(auto simp add: field_simps) - also have "\ = x$$i" apply(rule divide_eq_imp[OF Fal]) + unfolding dist_norm apply(subst euclidean_eq_iff) apply rule defer apply(rule, rule divide_nonneg_pos) prefer 4 + proof(rule) fix i :: 'a assume i:"i\Basis" + have "((1 - norm (a - x) / norm (a - b)) *\<^sub>R a + (norm (a - x) / norm (a - b)) *\<^sub>R b) \ i = + ((norm (a - b) - norm (a - x)) * (a \ i) + norm (a - x) * (b \ i)) / norm (a - b)" + using Fal by(auto simp add: field_simps inner_simps) + also have "\ = x\i" apply(rule divide_eq_imp[OF Fal]) unfolding as[unfolded dist_norm] using as[unfolded dist_triangle_eq] apply- - apply(subst (asm) euclidean_eq) using i apply(erule_tac x=i in allE) by(auto simp add:field_simps) - finally show "x $$ i = ((1 - norm (a - x) / norm (a - b)) *\<^sub>R a + (norm (a - x) / norm (a - b)) *\<^sub>R b) $$ i" + apply(subst (asm) euclidean_eq_iff) using i apply(erule_tac x=i in ballE) by(auto simp add:field_simps inner_simps) + finally show "x \ i = ((1 - norm (a - x) / norm (a - b)) *\<^sub>R a + (norm (a - x) / norm (a - b)) *\<^sub>R b) \ i" by auto - qed(insert Fal2, auto) qed qed + qed(insert Fal2, auto) qed +qed lemma between_midpoint: fixes a::"'a::euclidean_space" shows "between (a,b) (midpoint a b)" (is ?t1) "between (b,a) (midpoint a b)" (is ?t2) proof- have *:"\x y z. x = (1/2::real) *\<^sub>R z \ y = (1/2) *\<^sub>R z \ norm z = norm x + norm y" by auto show ?t1 ?t2 unfolding between midpoint_def dist_norm apply(rule_tac[!] *) - unfolding euclidean_eq[where 'a='a] - by(auto simp add:field_simps) qed + unfolding euclidean_eq_iff[where 'a='a] + by(auto simp add:field_simps inner_simps) qed lemma between_mem_convex_hull: "between (a,b) x \ x \ convex hull {a,b}" @@ -4303,7 +4308,7 @@ have *:"y = (1 - (1 - e)) *\<^sub>R ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) + (1 - e) *\<^sub>R x" using `e>0` by (auto simp add: scaleR_left_diff_distrib scaleR_right_diff_distrib) have "dist c ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) = abs(1/e) * norm (e *\<^sub>R c - y + (1 - e) *\<^sub>R x)" unfolding dist_norm unfolding norm_scaleR[symmetric] apply(rule arg_cong[where f=norm]) using `e>0` - by(auto simp add: euclidean_eq[where 'a='a] field_simps) + by(auto simp add: euclidean_eq_iff[where 'a='a] field_simps inner_simps) also have "\ = abs(1/e) * norm (x - e *\<^sub>R (x - c) - y)" by(auto intro!:arg_cong[where f=norm] simp add: algebra_simps) also have "\ < d" using as[unfolded dist_norm] and `e>0` by(auto simp add:pos_divide_less_eq[OF `e>0`] mult_commute) @@ -4347,236 +4352,237 @@ apply(rule_tac x=u in exI) defer apply(rule_tac x="\x. if x = 0 then 1 - setsum u s else u x" in exI) using assms(2) unfolding if_smult and setsum_delta_notmem[OF assms(2)] by auto -lemma substd_simplex: assumes "d\{.. x$$i = 0)}" +lemma substd_simplex: + assumes d: "d \ Basis" + shows "convex hull (insert 0 d) = {x. (\i\Basis. 0 \ x\i) \ (\i\d. x\i) \ 1 \ (\i\Basis. i \ d --> x\i = 0)}" (is "convex hull (insert 0 ?p) = ?s") -(* Proof is a modified copy of the proof of similar lemma std_simplex in Convex_Euclidean_Space.thy *) -proof- let ?D = d (*"{.. ?D} = basis ` ?D" by auto - note sumbas = this setsum_reindex[OF basis_inj_on[of d], unfolded o_def, OF assms] - show ?thesis unfolding simplex[OF finite_substdbasis `0 ~: ?p`] + from d have "finite d" by (blast intro: finite_subset finite_Basis) + show ?thesis unfolding simplex[OF `finite d` `0 ~: ?p`] apply(rule set_eqI) unfolding mem_Collect_eq apply rule apply(erule exE, (erule conjE)+) apply(erule_tac[2] conjE)+ proof- - fix x::"'a::euclidean_space" and u assume as: "\x\{basis i |i. i \?D}. 0 \ u x" - "setsum u {basis i |i. i \ ?D} \ 1" "(\x\{basis i |i. i \?D}. u x *\<^sub>R x) = x" - have *:"\i u (basis i) = x$$i" and "(!i x $$ i = 0)" using as(3) - unfolding sumbas unfolding substdbasis_expansion_unique[OF assms] by auto - hence **:"setsum u {basis i |i. i \ ?D} = setsum (op $$ x) ?D" unfolding sumbas + fix x::"'a::euclidean_space" and u assume as: "\x\?D. 0 \ u x" + "setsum u ?D \ 1" "(\x\?D. u x *\<^sub>R x) = x" + have *:"\i\Basis. i:d --> u i = x\i" and "(\i\Basis. i ~: d --> x \ i = 0)" using as(3) + unfolding substdbasis_expansion_unique[OF assms] by auto + hence **:"setsum u ?D = setsum (op \ x) ?D" apply-apply(rule setsum_cong2) using assms by auto - have " (\i x$$i) \ setsum (op $$ x) ?D \ 1" - apply - proof(rule,rule,rule) - fix i assume i:"i 0 \ x$$i" unfolding *[rule_format,OF i,symmetric] + have " (\i\Basis. 0 \ x\i) \ setsum (op \ x) ?D \ 1" + apply - proof(rule,rule) + fix i :: 'a assume i:"i\Basis" have "i : d ==> 0 \ x\i" unfolding *[rule_format,OF i,symmetric] apply(rule_tac as(1)[rule_format]) by auto - moreover have "i ~: d ==> 0 \ x$$i" - using `(!i x $$ i = 0)`[rule_format, OF i] by auto - ultimately show "0 \ x$$i" by auto + moreover have "i ~: d ==> 0 \ x\i" + using `(\i\Basis. i ~: d --> x \ i = 0)`[rule_format, OF i] by auto + ultimately show "0 \ x\i" by auto qed(insert as(2)[unfolded **], auto) - from this show " (\i x$$i) \ setsum (op $$ x) ?D \ 1 & (!i x $$ i = 0)" - using `(!i x $$ i = 0)` by auto - next fix x::"'a::euclidean_space" assume as:"\i x $$ i" "setsum (op $$ x) ?D \ 1" - "(!i x $$ i = 0)" - show "\u. (\x\{basis i |i. i \ ?D}. 0 \ u x) \ - setsum u {basis i |i. i \ ?D} \ 1 \ (\x\{basis i |i. i \ ?D}. u x *\<^sub>R x) = x" - apply(rule_tac x="\y. inner y x" in exI) apply(rule,rule) unfolding mem_Collect_eq apply(erule exE) - using as(1) apply(erule_tac x=i in allE) unfolding sumbas apply safe unfolding not_less basis_zero - unfolding substdbasis_expansion_unique[OF assms] euclidean_component_def[symmetric] - using as(2,3) by(auto simp add:dot_basis not_less) - qed qed + from this show " (\i\Basis. 0 \ x\i) \ setsum (op \ x) ?D \ 1 & (\i\Basis. i ~: d --> x \ i = 0)" + using `(\i\Basis. i ~: d --> x \ i = 0)` by auto + next fix x::"'a::euclidean_space" assume as:"\i\Basis. 0 \ x \ i" "setsum (op \ x) ?D \ 1" + "(\i\Basis. i ~: d --> x \ i = 0)" + show "\u. (\x\?D. 0 \ u x) \ setsum u ?D \ 1 \ (\x\?D. u x *\<^sub>R x) = x" + using as d unfolding substdbasis_expansion_unique[OF assms] + by (rule_tac x="inner x" in exI) auto + qed +qed lemma std_simplex: - "convex hull (insert 0 { basis i | i. ii x$$i) \ setsum (\i. x$$i) {.. 1 }" - using substd_simplex[of "{..i\Basis. 0 \ x\i) \ setsum (\i. x\i) Basis \ 1 }" + using substd_simplex[of Basis] by auto lemma interior_std_simplex: - "interior (convex hull (insert 0 { basis i| i. ii setsum (\i. x$$i) {..i\Basis. 0 < x\i) \ setsum (\i. x\i) Basis < 1 }" apply(rule set_eqI) unfolding mem_interior std_simplex unfolding subset_eq mem_Collect_eq Ball_def mem_ball unfolding Ball_def[symmetric] apply rule apply(erule exE, (erule conjE)+) defer apply(erule conjE) proof- - fix x::"'a" and e assume "0xa. dist x xa < e \ (\x xa $$ x) \ setsum (op $$ xa) {.. 1" - show "(\xa setsum (op $$ x) {..0` - unfolding dist_norm by (auto elim!:allE[where x=i]) - next have **:"dist x (x + (e / 2) *\<^sub>R basis 0) < e" using `e>0` - unfolding dist_norm by(auto intro!: mult_strict_left_mono) - have "\i. i (x + (e / 2) *\<^sub>R basis 0) $$ i = x$$i + (if i = 0 then e/2 else 0)" - by auto - hence *:"setsum (op $$ (x + (e / 2) *\<^sub>R basis 0)) {..i. x$$i + (if 0 = i then e/2 else 0)) {..xa. dist x xa < e \ (\x\Basis. 0 \ xa \ x) \ setsum (op \ xa) Basis \ 1" + show "(\xa\Basis. 0 < x \ xa) \ setsum (op \ x) Basis < 1" apply(safe) proof- + fix i :: 'a assume i:"i\Basis" thus "0 < x \ i" using as[THEN spec[where x="x - (e / 2) *\<^sub>R i"]] and `e>0` + unfolding dist_norm + by (auto elim!:ballE[where x=i] simp: inner_simps) + next have **:"dist x (x + (e / 2) *\<^sub>R (SOME i. i\Basis)) < e" using `e>0` + unfolding dist_norm by(auto intro!: mult_strict_left_mono simp: SOME_Basis) + have "\i. i\Basis \ (x + (e / 2) *\<^sub>R (SOME i. i\Basis)) \ i = x\i + (if i = (SOME i. i\Basis) then e/2 else 0)" + by (auto simp: SOME_Basis inner_Basis inner_simps) + hence *:"setsum (op \ (x + (e / 2) *\<^sub>R (SOME i. i\Basis))) Basis = setsum (\i. x\i + (if (SOME i. i\Basis) = i then e/2 else 0)) Basis" apply(rule_tac setsum_cong) by auto - have "setsum (op $$ x) {..R basis 0)) {.. x) Basis < setsum (op \ (x + (e / 2) *\<^sub>R (SOME i. i\Basis))) Basis" unfolding * setsum_addf + using `0 \ 1" using ** apply(drule_tac as[rule_format]) by auto - finally show "setsum (op $$ x) {..i x) Basis < 1" by auto qed +next fix x::"'a" assume as:"\i\Basis. 0 < x \ i" "setsum (op \ x) Basis < 1" guess a using UNIV_witness[where 'a='b] .. - let ?d = "(1 - setsum (op $$ x) {.. 0" apply(rule Min_grI) using as(1) by auto - moreover have"?d > 0" apply(rule divide_pos_pos) using as(2) by(auto simp add: Suc_le_eq) - ultimately show "\e>0. \y. dist x y < e \ (\i y $$ i) \ setsum (op $$ y) {.. 1" - apply(rule_tac x="min (Min ((op $$ x) ` {.. setsum (\i. x$$i + ?d) {..{.. x $$ i + ?d" by auto qed + let ?d = "(1 - setsum (op \ x) Basis) / real (DIM('a))" + have "Min ((op \ x) ` Basis) > 0" apply(rule Min_grI) using as(1) by auto + moreover have"?d > 0" apply(rule divide_pos_pos) using as(2) by (auto simp add: Suc_le_eq DIM_positive) + ultimately show "\e>0. \y. dist x y < e \ (\i\Basis. 0 \ y \ i) \ setsum (op \ y) Basis \ 1" + apply(rule_tac x="min (Min ((op \ x) ` Basis)) ?D" in exI) apply rule defer apply(rule,rule) proof- + fix y assume y:"dist x y < min (Min (op \ x ` Basis)) ?d" + have "setsum (op \ y) Basis \ setsum (\i. x\i + ?d) Basis" proof(rule setsum_mono) + fix i :: 'a assume i: "i\Basis" hence "abs (y\i - x\i) < ?d" apply-apply(rule le_less_trans) + using Basis_le_norm[OF i, of "y - x"] + using y[unfolded min_less_iff_conj dist_norm, THEN conjunct2] by(auto simp add: norm_minus_commute inner_diff_left) + thus "y \ i \ x \ i + ?d" by auto qed also have "\ \ 1" unfolding setsum_addf setsum_constant real_eq_of_nat by(auto simp add: Suc_le_eq) - finally show "(\i y $$ i) \ setsum (op $$ y) {.. 1" - proof safe fix i assume i:"ii\Basis. 0 \ y \ i) \ setsum (op \ y) Basis \ 1" + proof safe fix i :: 'a assume i:"i\Basis" + have "norm (x - y) < x\i" apply(rule less_le_trans) apply(rule y[unfolded min_less_iff_conj dist_norm, THEN conjunct1]) using i by auto - thus "0 \ y$$i" using component_le_norm[of "x - y" i] and as(1)[rule_format, of i] by auto + thus "0 \ y\i" using Basis_le_norm[OF i, of "x - y"] and as(1)[rule_format, OF i] + by (auto simp: inner_simps) qed qed auto qed lemma interior_std_simplex_nonempty: obtains a::"'a::euclidean_space" where - "a \ interior(convex hull (insert 0 {basis i | i . ij. if i = j then inverse (2 * real DIM('a)) else 0) ?D"]) apply(rule setsum_cong2) - defer apply(subst setsum_delta') unfolding euclidean_component_def using i by(auto simp add:dot_basis) } + "a \ interior(convex hull (insert 0 Basis))" proof- + let ?D = "Basis :: 'a set" let ?a = "setsum (\b::'a. inverse (2 * real DIM('a)) *\<^sub>R b) Basis" + { fix i :: 'a assume i:"i\Basis" have "?a \ i = inverse (2 * real DIM('a))" + by (rule trans[of _ "setsum (\j. if i = j then inverse (2 * real DIM('a)) else 0) ?D"]) + (simp_all add: setsum_cases i) } note ** = this show ?thesis apply(rule that[of ?a]) unfolding interior_std_simplex mem_Collect_eq proof safe - fix i assume i:"ii. inverse (2 * real DIM('a))) ?D" apply(rule setsum_cong2, rule **) by auto + fix i :: 'a assume i:"i\Basis" show "0 < ?a \ i" unfolding **[OF i] by(auto simp add: Suc_le_eq DIM_positive) + next have "setsum (op \ ?a) ?D = setsum (\i. inverse (2 * real DIM('a))) ?D" apply(rule setsum_cong2, rule **) by auto also have "\ < 1" unfolding setsum_constant real_eq_of_nat divide_inverse[symmetric] by (auto simp add:field_simps) - finally show "setsum (op $$ ?a) ?D < 1" by auto qed qed - -lemma rel_interior_substd_simplex: assumes "d\{..i\d. 0 < x$$i) & setsum (%i. x$$i) d < 1 & (!i x$$i = 0)}" + finally show "setsum (op \ ?a) ?D < 1" by auto qed qed + +lemma rel_interior_substd_simplex: assumes d: "d\Basis" + shows "rel_interior (convex hull (insert 0 d)) = + {x::'a::euclidean_space. (\i\d. 0 < x\i) \ (\i\d. x\i) < 1 \ (\i\Basis. i ~: d --> x\i = 0)}" (is "rel_interior (convex hull (insert 0 ?p)) = ?s") (* Proof is a modified copy of the proof of similar lemma interior_std_simplex in Convex_Euclidean_Space.thy *) proof- have "finite d" apply(rule finite_subset) using assms by auto -{ assume "d={}" hence ?thesis using rel_interior_sing using euclidean_eq[of _ 0] by auto } +{ assume "d={}" hence ?thesis using rel_interior_sing using euclidean_eq_iff[of _ 0] by auto } moreover { assume "d~={}" -have h0: "affine hull (convex hull (insert 0 ?p))={x::'a::euclidean_space. (!i x$$i = 0)}" +have h0: "affine hull (convex hull (insert 0 ?p))={x::'a::euclidean_space. (\i\Basis. i ~: d --> x\i = 0)}" using affine_hull_convex_hull affine_hull_substd_basis assms by auto -have aux: "!x::'n::euclidean_space. !i. ((! i:d. 0 <= x$$i) & (!i. i ~: d --> x$$i = 0))--> 0 <= x$$i" by auto +have aux: "!!x::'a. \i\Basis. ((\i\d. 0 \ x\i) \ (\i\Basis. i \ d \ x\i = 0)) \ 0 \ x\i" + by auto { fix x::"'a::euclidean_space" assume x_def: "x : rel_interior (convex hull (insert 0 ?p))" from this obtain e where e0: "e>0" and - "ball x e Int {xa. (!i xa$$i = 0)} <= convex hull (insert 0 ?p)" + "ball x e Int {xa. (\i\Basis. i ~: d --> xa\i = 0)} <= convex hull (insert 0 ?p)" using mem_rel_interior_ball[of x "convex hull (insert 0 ?p)"] h0 by auto - hence as: "ALL xa. (dist x xa < e & (!i xa$$i = 0)) --> - (!i : d. 0 <= xa $$ i) & setsum (op $$ xa) d <= 1" + hence as: "ALL xa. (dist x xa < e & (\i\Basis. i ~: d --> xa\i = 0)) --> + (!i : d. 0 <= xa \ i) & setsum (op \ xa) d <= 1" unfolding ball_def unfolding substd_simplex[OF assms] using assms by auto - have x0: "(!i x$$i = 0)" + have x0: "(\i\Basis. i ~: d --> x\i = 0)" using x_def rel_interior_subset substd_simplex[OF assms] by auto - have "(!i : d. 0 < x $$ i) & setsum (op $$ x) d < 1 & (!i x$$i = 0)" apply(rule,rule) + have "(\i\d. 0 < x \ i) & setsum (op \ x) d < 1 & (\i\Basis. i ~: d --> x\i = 0)" apply(rule,rule) proof- - fix i::nat assume "i:d" - hence "\ia\d. 0 \ (x - (e / 2) *\<^sub>R basis i) $$ ia" apply-apply(rule as[rule_format,THEN conjunct1]) - unfolding dist_norm using assms `e>0` x0 by auto - thus "0 < x $$ i" apply(erule_tac x=i in ballE) using `e>0` `i\d` assms by auto + fix i::'a assume "i\d" + hence "\ia\d. 0 \ (x - (e / 2) *\<^sub>R i) \ ia" apply-apply(rule as[rule_format,THEN conjunct1]) + unfolding dist_norm using d `e>0` x0 by (auto simp: inner_simps inner_Basis) + thus "0 < x \ i" apply(erule_tac x=i in ballE) using `e>0` `i\d` d + by (auto simp: inner_simps inner_Basis) next obtain a where a:"a:d" using `d ~= {}` by auto - have **:"dist x (x + (e / 2) *\<^sub>R basis a) < e" - using `e>0` and Euclidean_Space.norm_basis[of a] + then have **:"dist x (x + (e / 2) *\<^sub>R a) < e" + using `e>0` norm_Basis[of a] d unfolding dist_norm by auto - have "\i. (x + (e / 2) *\<^sub>R basis a) $$ i = x$$i + (if i = a then e/2 else 0)" - unfolding euclidean_simps using a assms by auto - hence *:"setsum (op $$ (x + (e / 2) *\<^sub>R basis a)) d = - setsum (\i. x$$i + (if a = i then e/2 else 0)) d" by(rule_tac setsum_cong, auto) - have h1: "(ALL i (x + (e / 2) *\<^sub>R basis a) $$ i = 0)" - using as[THEN spec[where x="x + (e / 2) *\<^sub>R basis a"]] `a:d` using x0 - by(auto elim:allE[where x=a]) - have "setsum (op $$ x) d < setsum (op $$ (x + (e / 2) *\<^sub>R basis a)) d" unfolding * setsum_addf + have "\i. i\Basis \ (x + (e / 2) *\<^sub>R a) \ i = x\i + (if i = a then e/2 else 0)" + using a d by (auto simp: inner_simps inner_Basis) + hence *:"setsum (op \ (x + (e / 2) *\<^sub>R a)) d = + setsum (\i. x\i + (if a = i then e/2 else 0)) d" using d by (intro setsum_cong) auto + have "a \ Basis" using `a \ d` d by auto + then have h1: "(\i\Basis. i ~: d --> (x + (e / 2) *\<^sub>R a) \ i = 0)" + using x0 d `a\d` by (auto simp add: inner_add_left inner_Basis) + have "setsum (op \ x) d < setsum (op \ (x + (e / 2) *\<^sub>R a)) d" unfolding * setsum_addf using `0 \ 1" using ** h1 as[rule_format, of "x + (e / 2) *\<^sub>R basis a"] by auto - finally show "setsum (op $$ x) d < 1 & (!i x$$i = 0)" using x0 by auto + also have "\ \ 1" using ** h1 as[rule_format, of "x + (e / 2) *\<^sub>R a"] by auto + finally show "setsum (op \ x) d < 1 & (\i\Basis. i ~: d --> x\i = 0)" using x0 by auto qed } moreover { fix x::"'a::euclidean_space" assume as: "x : ?s" - have "!i. ((0 0<=x$$i)" by auto + have "!i. ((0i) | (0=x\i) --> 0<=x\i)" by auto moreover have "!i. (i:d) | (i ~: d)" by auto ultimately - have "!i. ( (ALL i:d. 0 < x$$i) & (ALL i. i ~: d --> x$$i = 0) ) --> 0 <= x$$i" by metis + have "!i. ( (ALL i:d. 0 < x\i) & (ALL i. i ~: d --> x\i = 0) ) --> 0 <= x\i" by metis hence h2: "x : convex hull (insert 0 ?p)" using as assms unfolding substd_simplex[OF assms] by fastforce obtain a where a:"a:d" using `d ~= {}` by auto - let ?d = "(1 - setsum (op $$ x) d) / real (card d)" + let ?d = "(1 - setsum (op \ x) d) / real (card d)" have "0 < card d" using `d ~={}` `finite d` by (simp add: card_gt_0_iff) - have "Min ((op $$ x) ` d) > 0" using as `d \ {}` `finite d` by (simp add: Min_grI) + have "Min ((op \ x) ` d) > 0" using as `d \ {}` `finite d` by (simp add: Min_grI) moreover have "?d > 0" apply(rule divide_pos_pos) using as using `0 < card d` by auto - ultimately have h3: "min (Min ((op $$ x) ` d)) ?d > 0" by auto + ultimately have h3: "min (Min ((op \ x) ` d)) ?d > 0" by auto have "x : rel_interior (convex hull (insert 0 ?p))" unfolding rel_interior_ball mem_Collect_eq h0 apply(rule,rule h2) unfolding substd_simplex[OF assms] - apply(rule_tac x="min (Min ((op $$ x) ` d)) ?d" in exI) apply(rule,rule h3) apply safe unfolding mem_ball - proof- fix y::'a assume y:"dist x y < min (Min (op $$ x ` d)) ?d" and y2:"(!i y$$i = 0)" - have "setsum (op $$ y) d \ setsum (\i. x$$i + ?d) d" proof(rule setsum_mono) - fix i assume i:"i\d" - have "abs (y$$i - x$$i) < ?d" apply(rule le_less_trans) using component_le_norm[of "y - x" i] + apply(rule_tac x="min (Min ((op \ x) ` d)) ?d" in exI) apply(rule,rule h3) apply safe unfolding mem_ball + proof- + fix y::'a assume y:"dist x y < min (Min (op \ x ` d)) ?d" and y2: "\i\Basis. i \ d \ y\i = 0" + have "setsum (op \ y) d \ setsum (\i. x\i + ?d) d" + proof(rule setsum_mono) + fix i assume "i \ d" + with d have i: "i \ Basis" by auto + have "abs (y\i - x\i) < ?d" apply(rule le_less_trans) using Basis_le_norm[OF i, of "y - x"] using y[unfolded min_less_iff_conj dist_norm, THEN conjunct2] - by(auto simp add: norm_minus_commute) - thus "y $$ i \ x $$ i + ?d" by auto qed + by (auto simp add: norm_minus_commute inner_simps) + thus "y \ i \ x \ i + ?d" by auto + qed also have "\ \ 1" unfolding setsum_addf setsum_constant real_eq_of_nat using `0 < card d` by auto - finally show "setsum (op $$ y) d \ 1" . - - fix i assume "i y$$i" + finally show "setsum (op \ y) d \ 1" . + + fix i :: 'a assume i: "i \ Basis" thus "0 \ y\i" proof(cases "i\d") case True - have "norm (x - y) < x$$i" using y[unfolded min_less_iff_conj dist_norm, THEN conjunct1] - using Min_gr_iff[of "op $$ x ` d" "norm (x - y)"] `0 < card d` `i:d` + have "norm (x - y) < x\i" using y[unfolded min_less_iff_conj dist_norm, THEN conjunct1] + using Min_gr_iff[of "op \ x ` d" "norm (x - y)"] `0 < card d` `i:d` by (simp add: card_gt_0_iff) - thus "0 \ y$$i" using component_le_norm[of "x - y" i] and as(1)[rule_format] by auto + thus "0 \ y\i" using Basis_le_norm[OF i, of "x - y"] and as(1)[rule_format] + by (auto simp: inner_simps) qed(insert y2, auto) qed } ultimately have - "!!x :: 'a::euclidean_space. (x : rel_interior (convex hull insert 0 {basis i |i. i : d})) = - (x : {x. (ALL i:d. 0 < x $$ i) & - setsum (op $$ x) d < 1 & (ALL i x $$ i = 0)})" by blast + "\x. (x : rel_interior (convex hull insert 0 d)) = (x \ {x. (ALL i:d. 0 < x \ i) & + setsum (op \ x) d < 1 & (\i\Basis. i ~: d --> x \ i = 0)})" by blast from this have ?thesis by (rule set_eqI) } ultimately show ?thesis by blast qed -lemma rel_interior_substd_simplex_nonempty: assumes "d ~={}" "d\{..Basis" obtains a::"'a::euclidean_space" where - "a : rel_interior(convex hull (insert 0 {basis i | i . i : d}))" proof- + "a : rel_interior(convex hull (insert 0 d))" proof- (* Proof is a modified copy of the proof of similar lemma interior_std_simplex_nonempty in Convex_Euclidean_Space.thy *) - let ?D = d let ?a = "setsum (\b::'a::euclidean_space. inverse (2 * real (card d)) *\<^sub>R b) {(basis i) | i. i \ ?D}" - have *:"{basis i :: 'a | i. i \ ?D} = basis ` ?D" by auto + let ?D = d let ?a = "setsum (\b::'a::euclidean_space. inverse (2 * real (card d)) *\<^sub>R b) ?D" have "finite d" apply(rule finite_subset) using assms(2) by auto hence d1: "0 < real(card d)" using `d ~={}` by auto - { fix i assume "i:d" have "?a $$ i = inverse (2 * real (card d))" - unfolding * setsum_reindex[OF basis_inj_on, OF assms(2)] o_def + { fix i assume "i:d" + have "?a \ i = inverse (2 * real (card d))" apply(rule trans[of _ "setsum (\j. if i = j then inverse (2 * real (card d)) else 0) ?D"]) - unfolding euclidean_component_setsum + unfolding inner_setsum_left apply(rule setsum_cong2) using `i:d` `finite d` setsum_delta'[of d i "(%k. inverse (2 * real (card d)))"] d1 assms(2) - by (auto simp add: Euclidean_Space.basis_component[of i])} + by (auto simp: inner_simps inner_Basis set_rev_mp[OF _ assms(2)]) } note ** = this show ?thesis apply(rule that[of ?a]) unfolding rel_interior_substd_simplex[OF assms(2)] mem_Collect_eq proof safe fix i assume "i:d" have "0 < inverse (2 * real (card d))" using d1 by auto - also have "...=?a $$ i" using **[of i] `i:d` by auto - finally show "0 < ?a $$ i" by auto - next have "setsum (op $$ ?a) ?D = setsum (\i. inverse (2 * real (card d))) ?D" + also have "...=?a \ i" using **[of i] `i:d` by auto + finally show "0 < ?a \ i" by auto + next have "setsum (op \ ?a) ?D = setsum (\i. inverse (2 * real (card d))) ?D" by(rule setsum_cong2, rule **) also have "\ < 1" unfolding setsum_constant real_eq_of_nat divide_real_def[symmetric] by (auto simp add:field_simps) - finally show "setsum (op $$ ?a) ?D < 1" by auto - next fix i assume "iR (2 * real (card d)))" "{basis i |i. i : d}"]) - using finite_substdbasis[of d] apply blast + finally show "setsum (op \ ?a) ?D < 1" by auto + next fix i assume "i\Basis" and "i~:d" + have "?a : (span d)" + apply (rule span_setsum[of d "(%b. b /\<^sub>R (2 * real (card d)))" d]) + using finite_subset[OF assms(2) finite_Basis] + apply blast proof- - { fix x assume "(x :: 'a::euclidean_space): {basis i |i. i : d}" - hence "x : span {basis i |i. i : d}" - using span_superset[of _ "{basis i |i. i : d}"] by auto - hence "(x /\<^sub>R (2 * real (card d))) : (span {basis i |i. i : d})" - using span_mul[of x "{basis i |i. i : d}" "(inverse (real (card d)) / 2)"] by auto - } thus "\x\{basis i |i. i \ d}. x /\<^sub>R (2 * real (card d)) \ span {basis i ::'a |i. i \ d}" by auto + { fix x assume "(x :: 'a::euclidean_space): d" + hence "x : span d" + using span_superset[of _ "d"] by auto + hence "(x /\<^sub>R (2 * real (card d))) : (span d)" + using span_mul[of x "d" "(inverse (real (card d)) / 2)"] by auto + } thus "\x\d. x /\<^sub>R (2 * real (card d)) \ span d" by auto qed - thus "?a $$ i = 0 " using `i~:d` unfolding span_substd_basis[OF assms(2)] using `i i = 0 " using `i~:d` unfolding span_substd_basis[OF assms(2)] using `i\Basis` by auto qed qed @@ -4608,14 +4614,14 @@ ultimately have **: "affine hull (convex hull insert 0 B) = affine hull S" using affine_hull_span_0[of "convex hull insert 0 B"] affine_hull_span_0[of "S"] assms hull_subset[of S] by auto -obtain d and f::"'n=>'n" where fd: "card d = card B & linear f & f ` B = {basis i |i. i : (d :: nat set)} & - f ` span B = {x. ALL i x $$ i = (0::real)} & inj_on f (span B)" and d:"d\{..'n" where fd: "card d = card B & linear f & f ` B = d & + f ` span B = {x. \i\Basis. i ~: d --> x \ i = (0::real)} & inj_on f (span B)" and d:"d\Basis" using basis_to_substdbasis_subspace_isomorphism[of B,OF _ ] B_def by auto hence "bounded_linear f" using linear_conv_bounded_linear by auto have "d ~={}" using fd B_def `B ~={}` by auto -have "(insert 0 {basis i |i. i : d}) = f ` (insert 0 B)" using fd linear_0 by auto -hence "(convex hull (insert 0 {basis i |i. i : d})) = f ` (convex hull (insert 0 B))" - using convex_hull_linear_image[of f "(insert 0 {basis i |i. i : d})"] +have "(insert 0 d) = f ` (insert 0 B)" using fd linear_0 by auto +hence "(convex hull (insert 0 d)) = f ` (convex hull (insert 0 B))" + using convex_hull_linear_image[of f "(insert 0 d)"] convex_hull_linear_image[of f "(insert 0 B)"] `bounded_linear f` by auto moreover have "rel_interior (f ` (convex hull insert 0 B)) = f ` rel_interior (convex hull insert 0 B)" diff -r 46be26e02456 -r 899c9c4e4a4c src/HOL/Multivariate_Analysis/Derivative.thy --- a/src/HOL/Multivariate_Analysis/Derivative.thy Fri Dec 14 14:46:01 2012 +0100 +++ b/src/HOL/Multivariate_Analysis/Derivative.thy Fri Dec 14 15:46:01 2012 +0100 @@ -158,9 +158,6 @@ lemmas mult_left_has_derivative = bounded_linear.has_derivative [OF bounded_linear_mult_left] -lemmas euclidean_component_has_derivative = - bounded_linear.has_derivative [OF bounded_linear_euclidean_component] - lemma has_derivative_neg: assumes "(f has_derivative f') net" shows "((\x. -(f x)) has_derivative (\h. -(f' h))) net" @@ -191,20 +188,12 @@ using assms by (induct, simp_all add: has_derivative_const has_derivative_add) text {* Somewhat different results for derivative of scalar multiplier. *} -(** move **) -lemma linear_vmul_component: (* TODO: delete *) - assumes lf: "linear f" - shows "linear (\x. f x $$ k *\<^sub>R v)" - using lf - by (auto simp add: linear_def algebra_simps) - lemmas has_derivative_intros = has_derivative_id has_derivative_const has_derivative_add has_derivative_sub has_derivative_neg has_derivative_add_const scaleR_left_has_derivative scaleR_right_has_derivative inner_left_has_derivative inner_right_has_derivative - euclidean_component_has_derivative subsubsection {* Limit transformation for derivatives *} @@ -531,7 +520,7 @@ fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes "(f has_derivative f') (at x within s)" assumes "(f has_derivative f'') (at x within s)" - assumes "(\ie>0. \d. 0 < abs(d) \ abs(d) < e \ (x + d *\<^sub>R basis i) \ s)" + assumes "(\i\Basis. \e>0. \d. 0 < abs(d) \ abs(d) < e \ (x + d *\<^sub>R i) \ s)" shows "f' = f''" proof- note as = assms(1,2)[unfolded has_derivative_def] @@ -540,32 +529,32 @@ have "x islimpt s" unfolding islimpt_approachable proof(rule,rule) fix e::real assume "00`] .. + using assms(3)[rule_format,OF SOME_Basis `e>0`] .. thus "\x'\s. x' \ x \ dist x' x < e" - apply(rule_tac x="x + d *\<^sub>R basis 0" in bexI) - unfolding dist_norm by auto + apply(rule_tac x="x + d *\<^sub>R (SOME i. i \ Basis)" in bexI) + unfolding dist_norm by (auto simp: SOME_Basis nonzero_Basis) qed hence *:"netlimit (at x within s) = x" apply-apply(rule netlimit_within) unfolding trivial_limit_within by simp show ?thesis apply(rule linear_eq_stdbasis) unfolding linear_conv_bounded_linear apply(rule as(1,2)[THEN conjunct1])+ - proof(rule,rule,rule ccontr) - fix i assume i:"i "norm (f' (basis i) - f'' (basis i))" - assume "f' (basis i) \ f'' (basis i)" + proof(rule,rule ccontr) + fix i :: 'a assume i:"i \ Basis" def e \ "norm (f' i - f'' i)" + assume "f' i \ f'' i" hence "e>0" unfolding e_def by auto guess d using tendsto_diff [OF as(1,2)[THEN conjunct2], unfolded * Lim_within,rule_format,OF `e>0`] .. note d=this guess c using assms(3)[rule_format,OF i d[THEN conjunct1]] .. note c=this - have *:"norm (- ((1 / \c\) *\<^sub>R f' (c *\<^sub>R basis i)) + (1 / \c\) *\<^sub>R f'' (c *\<^sub>R basis i)) = norm ((1 / abs c) *\<^sub>R (- (f' (c *\<^sub>R basis i)) + f'' (c *\<^sub>R basis i)))" + have *:"norm (- ((1 / \c\) *\<^sub>R f' (c *\<^sub>R i)) + (1 / \c\) *\<^sub>R f'' (c *\<^sub>R i)) = norm ((1 / abs c) *\<^sub>R (- (f' (c *\<^sub>R i)) + f'' (c *\<^sub>R i)))" unfolding scaleR_right_distrib by auto - also have "\ = norm ((1 / abs c) *\<^sub>R (c *\<^sub>R (- (f' (basis i)) + f'' (basis i))))" + also have "\ = norm ((1 / abs c) *\<^sub>R (c *\<^sub>R (- (f' i) + f'' i)))" unfolding f'.scaleR f''.scaleR unfolding scaleR_right_distrib scaleR_minus_right by auto also have "\ = e" unfolding e_def using c[THEN conjunct1] - using norm_minus_cancel[of "f' (basis i) - f'' (basis i)"] + using norm_minus_cancel[of "f' i - f'' i"] by (auto simp add: add.commute ab_diff_minus) finally show False using c - using d[THEN conjunct2,rule_format,of "x + c *\<^sub>R basis i"] + using d[THEN conjunct2,rule_format,of "x + c *\<^sub>R i"] unfolding dist_norm unfolding f'.scaleR f''.scaleR f'.add f''.add f'.diff f''.diff scaleR_scaleR scaleR_right_diff_distrib scaleR_right_distrib @@ -584,37 +573,38 @@ lemma frechet_derivative_unique_within_closed_interval: fixes f::"'a::ordered_euclidean_space \ 'b::real_normed_vector" - assumes "\i {a..b}" (is "x\?I") + assumes "\i\Basis. a\i < b\i" "x \ {a..b}" (is "x\?I") assumes "(f has_derivative f' ) (at x within {a..b})" assumes "(f has_derivative f'') (at x within {a..b})" shows "f' = f''" apply(rule frechet_derivative_unique_within) apply(rule assms(3,4))+ -proof(rule,rule,rule,rule) - fix e::real and i assume "e>0" and i:"id. 0 < \d\ \ \d\ < e \ x + d *\<^sub>R basis i \ {a..b}" - proof(cases "x$$i=a$$i") +proof(rule,rule,rule) + fix e::real and i :: 'a assume "e>0" and i:"i\Basis" + thus "\d. 0 < \d\ \ \d\ < e \ x + d *\<^sub>R i \ {a..b}" + proof(cases "x\i=a\i") case True thus ?thesis - apply(rule_tac x="(min (b$$i - a$$i) e) / 2" in exI) - using assms(1)[THEN spec[where x=i]] and `e>0` and assms(2) - unfolding mem_interval euclidean_simps - using i by (auto simp add: field_simps) - next note * = assms(2)[unfolded mem_interval,THEN spec[where x=i]] - case False moreover have "a $$ i < x $$ i" using False * by auto + apply(rule_tac x="(min (b\i - a\i) e) / 2" in exI) + using assms(1)[THEN bspec[where x=i]] and `e>0` and assms(2) + unfolding mem_interval + using i by (auto simp add: field_simps inner_simps inner_Basis) + next + note * = assms(2)[unfolded mem_interval, THEN bspec, OF i] + case False moreover have "a \ i < x \ i" using False * by auto moreover { - have "a $$ i * 2 + min (x $$ i - a $$ i) e \ a$$i *2 + x$$i - a$$i" + have "a \ i * 2 + min (x \ i - a \ i) e \ a\i *2 + x\i - a\i" by auto - also have "\ = a$$i + x$$i" by auto - also have "\ \ 2 * x$$i" using * by auto - finally have "a $$ i * 2 + min (x $$ i - a $$ i) e \ x $$ i * 2" by auto + also have "\ = a\i + x\i" by auto + also have "\ \ 2 * (x\i)" using * by auto + finally have "a \ i * 2 + min (x \ i - a \ i) e \ x \ i * 2" by auto } - moreover have "min (x $$ i - a $$ i) e \ 0" using * and `e>0` by auto - hence "x $$ i * 2 \ b $$ i * 2 + min (x $$ i - a $$ i) e" using * by auto + moreover have "min (x \ i - a \ i) e \ 0" using * and `e>0` by auto + hence "x \ i * 2 \ b \ i * 2 + min (x \ i - a \ i) e" using * by auto ultimately show ?thesis - apply(rule_tac x="- (min (x$$i - a$$i) e) / 2" in exI) - using assms(1)[THEN spec[where x=i]] and `e>0` and assms(2) - unfolding mem_interval euclidean_simps - using i by (auto simp add: field_simps) + apply(rule_tac x="- (min (x\i - a\i) e) / 2" in exI) + using assms(1)[THEN bspec, OF i] and `e>0` and assms(2) + unfolding mem_interval + using i by (auto simp add: field_simps inner_simps inner_Basis) qed qed @@ -638,7 +628,7 @@ lemma frechet_derivative_within_closed_interval: fixes f::"'a::ordered_euclidean_space \ 'b::real_normed_vector" - assumes "\i {a..b}" + assumes "\i\Basis. a\i < b\i" and "x \ {a..b}" assumes "(f has_derivative f') (at x within {a.. b})" shows "frechet_derivative f (at x within {a.. b}) = f'" apply(rule frechet_derivative_unique_within_closed_interval[where f=f]) @@ -650,14 +640,14 @@ lemma linear_componentwise: fixes f:: "'a::euclidean_space \ 'b::euclidean_space" assumes lf: "linear f" - shows "(f x) $$ j = (\i j = (\i\Basis. (x\i) * (f i\j))" (is "?lhs = ?rhs") proof - - have fA: "finite {..iR f (basis i))$$j" - by simp + have fA: "finite Basis" by simp + have "?rhs = (\i\Basis. (x\i) *\<^sub>R (f i))\j" + by (simp add: inner_setsum_left) then show ?thesis unfolding linear_setsum_mul[OF lf fA, symmetric] - unfolding euclidean_representation[symmetric] .. + unfolding euclidean_representation .. qed text {* We do not introduce @{text jacobian}, which is defined on matrices, instead we use @@ -665,52 +655,54 @@ lemma jacobian_works: "(f::('a::euclidean_space) \ ('b::euclidean_space)) differentiable net \ - (f has_derivative (\h. \\ i. - \j (f has_derivative (\h. \\ i. ?SUM h i)) net") + (f has_derivative (\h. \i\Basis. + (\j\Basis. frechet_derivative f net (j) \ i * (h \ j)) *\<^sub>R i)) net" + (is "?differentiable \ (f has_derivative (\h. \i\Basis. ?SUM h i *\<^sub>R i)) net") proof assume *: ?differentiable { fix h i - have "?SUM h i = frechet_derivative f net h $$ i" using * + have "?SUM h i = frechet_derivative f net h \ i" using * by (auto intro!: setsum_cong simp: linear_componentwise[of _ h i] linear_frechet_derivative) } - thus "(f has_derivative (\h. \\ i. ?SUM h i)) net" - using * by (simp add: frechet_derivative_works) + with * show "(f has_derivative (\h. \i\Basis. ?SUM h i *\<^sub>R i)) net" + by (simp add: frechet_derivative_works euclidean_representation) qed (auto intro!: differentiableI) lemma differential_zero_maxmin_component: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" - assumes k: "k < DIM('b)" - and ball: "0 < e" "((\y \ ball x e. (f y)$$k \ (f x)$$k) \ (\y\ball x e. (f x)$$k \ (f y)$$k))" + assumes k: "k \ Basis" + and ball: "0 < e" "((\y \ ball x e. (f y)\k \ (f x)\k) \ (\y\ball x e. (f x)\k \ (f y)\k))" and diff: "f differentiable (at x)" - shows "(\\ j. frechet_derivative f (at x) (basis j) $$ k) = (0::'a)" (is "?D k = 0") + shows "(\j\Basis. (frechet_derivative f (at x) j \ k) *\<^sub>R j) = (0::'a)" (is "?D k = 0") proof (rule ccontr) assume "?D k \ 0" - then obtain j where j: "?D k $$ j \ 0" "j < DIM('a)" - unfolding euclidean_lambda_beta euclidean_eq[of _ "0::'a"] by auto - hence *: "\?D k $$ j\ / 2 > 0" by auto + then obtain j where j: "?D k \ j \ 0" "j \ Basis" + unfolding euclidean_eq_iff[of _ "0::'a"] by auto + hence *: "\?D k \ j\ / 2 > 0" by auto note as = diff[unfolded jacobian_works has_derivative_at_alt] guess e' using as[THEN conjunct2, rule_format, OF *] .. note e' = this guess d using real_lbound_gt_zero[OF ball(1) e'[THEN conjunct1]] .. note d = this { fix c assume "abs c \ d" - hence *:"norm (x + c *\<^sub>R basis j - x) < e'" using norm_basis[of j] d by auto - let ?v = "(\\ i. \lR basis j :: 'a) $$ l)" + hence *:"norm (x + c *\<^sub>R j - x) < e'" using norm_Basis[OF j(2)] d by auto + let ?v = "(\i\Basis. (\l\Basis. ?D i \ l * ((c *\<^sub>R j :: 'a) \ l)) *\<^sub>R i)" have if_dist: "\ P a b c. a * (if P then b else c) = (if P then a * b else a * c)" by auto - have "\(f (x + c *\<^sub>R basis j) - f x - ?v) $$ k\ \ - norm (f (x + c *\<^sub>R basis j) - f x - ?v)" by (rule component_le_norm) - also have "\ \ \?D k $$ j\ / 2 * \c\" - using e'[THEN conjunct2, rule_format, OF *] and norm_basis[of j] by fastforce - finally have "\(f (x + c *\<^sub>R basis j) - f x - ?v) $$ k\ \ \?D k $$ j\ / 2 * \c\" by simp - hence "\f (x + c *\<^sub>R basis j) $$ k - f x $$ k - c * ?D k $$ j\ \ \?D k $$ j\ / 2 * \c\" - unfolding euclidean_simps euclidean_lambda_beta using j k - by (simp add: if_dist setsum_cases field_simps) } note * = this - have "x + d *\<^sub>R basis j \ ball x e" "x - d *\<^sub>R basis j \ ball x e" - unfolding mem_ball dist_norm using norm_basis[of j] d by auto - hence **:"((f (x - d *\<^sub>R basis j))$$k \ (f x)$$k \ (f (x + d *\<^sub>R basis j))$$k \ (f x)$$k) \ - ((f (x - d *\<^sub>R basis j))$$k \ (f x)$$k \ (f (x + d *\<^sub>R basis j))$$k \ (f x)$$k)" using ball by auto + have "\(f (x + c *\<^sub>R j) - f x - ?v) \ k\ \ + norm (f (x + c *\<^sub>R j) - f x - ?v)" by (rule Basis_le_norm[OF k]) + also have "\ \ \?D k \ j\ / 2 * \c\" + using e'[THEN conjunct2, rule_format, OF *] and norm_Basis[OF j(2)] j + by simp + finally have "\(f (x + c *\<^sub>R j) - f x - ?v) \ k\ \ \?D k \ j\ / 2 * \c\" by simp + hence "\f (x + c *\<^sub>R j) \ k - f x \ k - c * (?D k \ j)\ \ \?D k \ j\ / 2 * \c\" + using j k + by (simp add: inner_simps field_simps inner_Basis setsum_cases if_dist) } + note * = this + have "x + d *\<^sub>R j \ ball x e" "x - d *\<^sub>R j \ ball x e" + unfolding mem_ball dist_norm using norm_Basis[OF j(2)] d by auto + hence **:"((f (x - d *\<^sub>R j))\k \ (f x)\k \ (f (x + d *\<^sub>R j))\k \ (f x)\k) \ + ((f (x - d *\<^sub>R j))\k \ (f x)\k \ (f (x + d *\<^sub>R j))\k \ (f x)\k)" using ball by auto have ***: "\y y1 y2 d dx::real. (y1\y\y2\y) \ (y\y1\y\y2) \ d < abs dx \ abs(y1 - y - - dx) \ d \ (abs (y2 - y - dx) \ d) \ False" by arith - show False apply(rule ***[OF **, where dx="d * ?D k $$ j" and d="\?D k $$ j\ / 2 * \d\"]) + show False apply(rule ***[OF **, where dx="d * (?D k \ j)" and d="\?D k \ j\ / 2 * \d\"]) using *[of "-d"] and *[of d] and d[THEN conjunct1] and j unfolding mult_minus_left unfolding abs_mult diff_minus_eq_add scaleR_minus_left @@ -728,13 +720,13 @@ proof - obtain e where e:"e>0" "ball x e \ s" using `open s`[unfolded open_contains_ball] and `x \ s` by auto - with differential_zero_maxmin_component[where 'b=real, of 0 e x f, simplified] - have "(\\ j. frechet_derivative f (at x) (basis j)) = (0::'a)" - unfolding differentiable_def using mono deriv by auto + with differential_zero_maxmin_component[where 'b=real, of 1 e x f] mono deriv + have "(\j\Basis. frechet_derivative f (at x) j *\<^sub>R j) = (0::'a)" + by (auto simp: Basis_real_def differentiable_def) with frechet_derivative_at[OF deriv, symmetric] - have "\ii\Basis. f' i = 0" + by (simp add: euclidean_eq_iff[of _ "0::'a"] inner_setsum_left_Basis) + with derivative_is_linear[OF deriv, THEN linear_componentwise, of _ 1] show ?thesis by (simp add: fun_eq_iff) qed @@ -1281,8 +1273,8 @@ proof- interpret bounded_linear g' using assms by auto note f'g' = assms(4)[unfolded id_def o_def,THEN cong] - have "g' (f' a (\\ i.1)) = (\\ i.1)" "(\\ i.1) \ (0::'n)" defer - apply(subst euclidean_eq) using f'g' by auto + have "g' (f' a (\Basis)) = (\Basis)" "(\Basis) \ (0::'n)" defer + apply(subst euclidean_eq_iff) using f'g' by auto hence *:"0 < onorm g'" unfolding onorm_pos_lt[OF assms(3)[unfolded linear_linear]] by fastforce def k \ "1 / onorm g' / 2" have *:"k>0" unfolding k_def using * by auto @@ -1726,7 +1718,7 @@ have *:"(\x. x *\<^sub>R f') = (\x. x *\<^sub>R f'')" apply(rule frechet_derivative_unique_within_closed_interval[of "a" "b"]) using assms(3-)[unfolded has_vector_derivative_def] using assms(1-2) - by auto + by (auto simp: Basis_real_def) show ?thesis proof(rule ccontr) assume "f' \ f''" diff -r 46be26e02456 -r 899c9c4e4a4c src/HOL/Multivariate_Analysis/Determinants.thy --- a/src/HOL/Multivariate_Analysis/Determinants.thy Fri Dec 14 14:46:01 2012 +0100 +++ b/src/HOL/Multivariate_Analysis/Determinants.thy Fri Dec 14 15:46:01 2012 +0100 @@ -452,7 +452,7 @@ ultimately show ?thesis apply - - apply (rule span_induct_alt[of ?P ?S, OF P0, folded smult_conv_scaleR]) + apply (rule span_induct_alt[of ?P ?S, OF P0, folded scalar_mult_eq_scaleR]) apply blast apply (rule x) done @@ -746,7 +746,7 @@ apply (rule span_setsum) apply simp apply (rule ballI) - apply (rule span_mul [where 'a="real^'n", folded smult_conv_scaleR])+ + apply (rule span_mul [where 'a="real^'n", folded scalar_mult_eq_scaleR])+ apply (rule span_superset) apply auto done @@ -782,7 +782,7 @@ apply (rule det_row_span) apply (rule span_setsum[OF fUk]) apply (rule ballI) - apply (rule span_mul [where 'a="real^'n", folded smult_conv_scaleR])+ + apply (rule span_mul [where 'a="real^'n", folded scalar_mult_eq_scaleR])+ apply (rule span_superset) apply auto done @@ -879,9 +879,10 @@ have th0: "\b (x::'a::comm_ring_1). (if b then 1 else 0)*x = (if b then x else 0)" "\b (x::'a::comm_ring_1). x*(if b then 1 else 0) = (if b then x else 0)" by simp_all - from fd[rule_format, of "cart_basis i" "cart_basis j", unfolded matrix_works[OF lf, symmetric] dot_matrix_vector_mul] + from fd[rule_format, of "axis i 1" "axis j 1", unfolded matrix_works[OF lf, symmetric] dot_matrix_vector_mul] have "?A$i$j = ?m1 $ i $ j" - by (simp add: inner_vec_def matrix_matrix_mult_def columnvector_def rowvector_def cart_basis_def th0 setsum_delta[OF fU] mat_def)} + by (simp add: inner_vec_def matrix_matrix_mult_def columnvector_def rowvector_def + th0 setsum_delta[OF fU] mat_def axis_def) } hence "orthogonal_matrix ?mf" unfolding orthogonal_matrix by vector with lf have ?rhs by blast} moreover @@ -931,7 +932,9 @@ unfolding dot_norm_neg dist_norm[symmetric] unfolding th0 fd[rule_format] by (simp add: power2_eq_square field_simps)} note fc = this - show ?thesis unfolding linear_def vector_eq[where 'a="real^'n"] smult_conv_scaleR by (simp add: inner_add fc field_simps) + show ?thesis + unfolding linear_def vector_eq[where 'a="real^'n"] scalar_mult_eq_scaleR + by (simp add: inner_add fc field_simps) qed lemma isometry_linear: @@ -981,7 +984,7 @@ apply (subst H(4)) using H(5-9) apply (simp add: norm_eq norm_eq_1) - apply (simp add: inner_diff smult_conv_scaleR) unfolding * + apply (simp add: inner_diff scalar_mult_eq_scaleR) unfolding * by (simp add: field_simps) } note th0 = this let ?g = "\x. if x = 0 then 0 else norm x *\<^sub>R f (inverse (norm x) *\<^sub>R x)" diff -r 46be26e02456 -r 899c9c4e4a4c src/HOL/Multivariate_Analysis/Euclidean_Space.thy --- a/src/HOL/Multivariate_Analysis/Euclidean_Space.thy Fri Dec 14 14:46:01 2012 +0100 +++ b/src/HOL/Multivariate_Analysis/Euclidean_Space.thy Fri Dec 14 15:46:01 2012 +0100 @@ -23,24 +23,24 @@ assumes euclidean_all_zero_iff: "(\u\Basis. inner x u = 0) \ (x = 0)" - -- "FIXME: make this a separate definition" - fixes dimension :: "'a itself \ nat" - assumes dimension_def: "dimension TYPE('a) = card Basis" - - -- "FIXME: eventually basis function can be removed" - fixes basis :: "nat \ 'a" - assumes image_basis: "basis ` {.. nat" where + "dimension TYPE('a) \ card (Basis :: 'a set)" syntax "_type_dimension" :: "type => nat" ("(1DIM/(1'(_')))") translations "DIM('t)" == "CONST dimension (TYPE('t))" -lemma (in euclidean_space) norm_Basis: "u \ Basis \ norm u = 1" +lemma (in euclidean_space) norm_Basis[simp]: "u \ Basis \ norm u = 1" unfolding norm_eq_sqrt_inner by (simp add: inner_Basis) +lemma (in euclidean_space) inner_same_Basis[simp]: "u \ Basis \ inner u u = 1" + by (simp add: inner_Basis) + +lemma (in euclidean_space) inner_not_same_Basis: "u \ Basis \ v \ Basis \ u \ v \ inner u v = 0" + by (simp add: inner_Basis) + lemma (in euclidean_space) sgn_Basis: "u \ Basis \ sgn u = u" - unfolding sgn_div_norm by (simp add: norm_Basis scaleR_one) + unfolding sgn_div_norm by (simp add: scaleR_one) lemma (in euclidean_space) Basis_zero [simp]: "0 \ Basis" proof @@ -51,184 +51,45 @@ lemma (in euclidean_space) nonzero_Basis: "u \ Basis \ u \ 0" by clarsimp -text {* Lemmas related to @{text basis} function. *} - -lemma (in euclidean_space) euclidean_all_zero: - "(\i (x = 0)" - using euclidean_all_zero_iff [of x, folded image_basis] - unfolding ball_simps by (simp add: Ball_def inner_commute) - -lemma (in euclidean_space) basis_zero [simp]: - "DIM('a) \ i \ basis i = 0" - using basis_finite by auto +lemma (in euclidean_space) SOME_Basis: "(SOME i. i \ Basis) \ Basis" + by (metis ex_in_conv nonempty_Basis someI_ex) -lemma (in euclidean_space) DIM_positive [intro]: "0 < DIM('a)" - unfolding dimension_def by (simp add: card_gt_0_iff) - -lemma (in euclidean_space) basis_inj [simp, intro]: "inj_on basis {.. Basis \ i < DIM('a)" - by (cases "i < DIM('a)", simp add: image_basis [symmetric], simp) - -lemma (in euclidean_space) Basis_elim: - assumes "u \ Basis" obtains i where "i < DIM('a)" and "u = basis i" - using assms unfolding image_basis [symmetric] by fast +lemma (in euclidean_space) inner_setsum_left_Basis[simp]: + "b \ Basis \ inner (\i\Basis. f i *\<^sub>R i) b = f b" + by (simp add: inner_setsum_left inner_Basis if_distrib setsum_cases) -lemma (in euclidean_space) basis_orthonormal: - "\ij i < DIM('a) then 1 else 0)" -proof (cases "(i < DIM('a) \ j < DIM('a))") - case False - hence "inner (basis i) (basis j) = 0" by auto - thus ?thesis using False by auto -next - case True thus ?thesis using basis_orthonormal by auto -qed - -lemma (in euclidean_space) basis_eq_0_iff [simp]: - "basis i = 0 \ DIM('a) \ i" +lemma (in euclidean_space) euclidean_eqI: + assumes b: "\b. b \ Basis \ inner x b = inner y b" shows "x = y" proof - - have "inner (basis i) (basis i) = 0 \ DIM('a) \ i" - by (simp add: dot_basis) - thus ?thesis by simp + from b have "\b\Basis. inner (x - y) b = 0" + by (simp add: inner_diff_left) + then show "x = y" + by (simp add: euclidean_all_zero_iff) qed -lemma (in euclidean_space) norm_basis [simp]: - "norm (basis i) = (if i < DIM('a) then 1 else 0)" - unfolding norm_eq_sqrt_inner dot_basis by simp - -lemma (in euclidean_space) basis_neq_0 [intro]: - assumes "i 0" - using assms by simp - -subsubsection {* Projecting components *} - -definition (in euclidean_space) euclidean_component (infixl "$$" 90) - where "x $$ i = inner (basis i) x" - -lemma bounded_linear_euclidean_component: - "bounded_linear (\x. euclidean_component x i)" - unfolding euclidean_component_def - by (rule bounded_linear_inner_right) - -lemmas tendsto_euclidean_component [tendsto_intros] = - bounded_linear.tendsto [OF bounded_linear_euclidean_component] - -lemmas isCont_euclidean_component [simp] = - bounded_linear.isCont [OF bounded_linear_euclidean_component] - -lemma euclidean_component_zero [simp]: "0 $$ i = 0" - unfolding euclidean_component_def by (rule inner_zero_right) - -lemma euclidean_component_add [simp]: "(x + y) $$ i = x $$ i + y $$ i" - unfolding euclidean_component_def by (rule inner_add_right) - -lemma euclidean_component_diff [simp]: "(x - y) $$ i = x $$ i - y $$ i" - unfolding euclidean_component_def by (rule inner_diff_right) - -lemma euclidean_component_minus [simp]: "(- x) $$ i = - (x $$ i)" - unfolding euclidean_component_def by (rule inner_minus_right) - -lemma euclidean_component_scaleR [simp]: "(scaleR a x) $$ i = a * (x $$ i)" - unfolding euclidean_component_def by (rule inner_scaleR_right) - -lemma euclidean_component_setsum [simp]: "(\x\A. f x) $$ i = (\x\A. f x $$ i)" - unfolding euclidean_component_def by (rule inner_setsum_right) - -lemma euclidean_eqI: - fixes x y :: "'a::euclidean_space" - assumes "\i. i < DIM('a) \ x $$ i = y $$ i" shows "x = y" -proof - - from assms have "\i (\i (\b\Basis. inner x b = inner y b)" by (auto intro: euclidean_eqI) -lemma (in euclidean_space) basis_component [simp]: - "basis i $$ j = (if i = j \ i < DIM('a) then 1 else 0)" - unfolding euclidean_component_def dot_basis by auto - -lemma (in euclidean_space) basis_at_neq_0 [intro]: - "i < DIM('a) \ basis i $$ i \ 0" - by simp - -lemma (in euclidean_space) euclidean_component_ge [simp]: - assumes "i \ DIM('a)" shows "x $$ i = 0" - unfolding euclidean_component_def basis_zero[OF assms] by simp +lemma (in euclidean_space) euclidean_representation_setsum: + "(\i\Basis. f i *\<^sub>R i) = b \ (\i\Basis. f i = inner b i)" + by (subst euclidean_eq_iff) simp -lemmas euclidean_simps = - euclidean_component_add - euclidean_component_diff - euclidean_component_scaleR - euclidean_component_minus - euclidean_component_setsum - basis_component - -lemma euclidean_representation: - fixes x :: "'a::euclidean_space" - shows "x = (\iR basis i)" - apply (rule euclidean_eqI) - apply (simp add: if_distrib setsum_delta cong: if_cong) - done - -subsubsection {* Binder notation for vectors *} - -definition (in euclidean_space) Chi (binder "\\ " 10) where - "(\\ i. f i) = (\iR basis i)" +lemma (in euclidean_space) euclidean_representation: "(\b\Basis. inner x b *\<^sub>R b) = x" + unfolding euclidean_representation_setsum by simp -lemma euclidean_lambda_beta [simp]: - "((\\ i. f i)::'a::euclidean_space) $$ j = (if j < DIM('a) then f j else 0)" - by (auto simp: Chi_def if_distrib setsum_cases intro!: setsum_cong) - -lemma euclidean_lambda_beta': - "j < DIM('a) \ ((\\ i. f i)::'a::euclidean_space) $$ j = f j" - by simp - -lemma euclidean_lambda_beta'':"(\j < DIM('a::euclidean_space). P j (((\\ i. f i)::'a) $$ j)) \ - (\j < DIM('a::euclidean_space). P j (f j))" by auto - -lemma euclidean_beta_reduce[simp]: - "(\\ i. x $$ i) = (x::'a::euclidean_space)" - by (simp add: euclidean_eq) - -lemma euclidean_lambda_beta_0[simp]: - "((\\ i. f i)::'a::euclidean_space) $$ 0 = f 0" - by (simp add: DIM_positive) +lemma (in euclidean_space) choice_Basis_iff: + fixes P :: "'a \ real \ bool" + shows "(\i\Basis. \x. P i x) \ (\x. \i\Basis. P i (inner x i))" + unfolding bchoice_iff +proof safe + fix f assume "\i\Basis. P i (f i)" + then show "\x. \i\Basis. P i (inner x i)" + by (auto intro!: exI[of _ "\i\Basis. f i *\<^sub>R i"]) +qed auto -lemma euclidean_inner: - "inner x (y::'a) = (\ii. dist (x $$ i) (y $$ i)) {..x$$i\ \ norm (x::'a::euclidean_space)" - unfolding euclidean_component_def - by (rule order_trans [OF Cauchy_Schwarz_ineq2]) simp - -lemma dist_nth_le: "dist (x $$ i) (y $$ i) \ dist x (y::'a::euclidean_space)" - unfolding euclidean_dist_l2 [where 'a='a] - by (cases "i < DIM('a)", rule member_le_setL2, auto) +lemma DIM_positive: "0 < DIM('a::euclidean_space)" + by (simp add: card_gt_0_iff) subsection {* Subclass relationships *} @@ -239,11 +100,13 @@ assume "open {x}" then obtain e where "0 < e" and e: "\y. dist y x < e \ y = x" unfolding open_dist by fast - def y \ "x + scaleR (e/2) (sgn (basis 0))" + def y \ "x + scaleR (e/2) (SOME b. b \ Basis)" + have [simp]: "(SOME b. b \ Basis) \ Basis" + by (rule someI_ex) (auto simp: ex_in_conv) from `0 < e` have "y \ x" - unfolding y_def by (simp add: sgn_zero_iff DIM_positive) + unfolding y_def by (auto intro!: nonzero_Basis) from `0 < e` have "dist y x < e" - unfolding y_def by (simp add: dist_norm norm_sgn) + unfolding y_def by (simp add: dist_norm norm_Basis) from `y \ x` and `dist y x < e` show "False" using e by simp qed @@ -256,23 +119,17 @@ instantiation real :: euclidean_space begin -definition - "Basis = {1::real}" - -definition - "dimension (t::real itself) = 1" - -definition [simp]: - "basis i = (if i = 0 then 1 else (0::real))" - -lemma DIM_real [simp]: "DIM(real) = 1" - by (rule dimension_real_def) +definition + [simp]: "Basis = {1::real}" instance by default (auto simp add: Basis_real_def) end +lemma DIM_real[simp]: "DIM(real) = 1" + by simp + subsubsection {* Type @{typ complex} *} instantiation complex :: euclidean_space @@ -281,20 +138,13 @@ definition Basis_complex_def: "Basis = {1, ii}" -definition - "dimension (t::complex itself) = 2" - -definition - "basis i = (if i = 0 then 1 else if i = 1 then ii else 0)" - instance - by default (auto simp add: Basis_complex_def dimension_complex_def - basis_complex_def intro: complex_eqI split: split_if_asm) + by default (auto simp add: Basis_complex_def intro: complex_eqI split: split_if_asm) end lemma DIM_complex[simp]: "DIM(complex) = 2" - by (rule dimension_complex_def) + unfolding Basis_complex_def by simp subsubsection {* Type @{typ "'a \ 'b"} *} @@ -304,12 +154,6 @@ definition "Basis = (\u. (u, 0)) ` Basis \ (\v. (0, v)) ` Basis" -definition - "dimension (t::('a \ 'b) itself) = DIM('a) + DIM('b)" - -definition - "basis i = (if i < DIM('a) then (basis i, 0) else (0, basis (i - DIM('a))))" - instance proof show "(Basis :: ('a \ 'b) set) \ {}" unfolding Basis_prod_def by simp @@ -327,20 +171,12 @@ show "(\u\Basis. inner x u = 0) \ x = 0" unfolding Basis_prod_def ball_Un ball_simps by (simp add: inner_prod_def prod_eq_iff euclidean_all_zero_iff) -next - show "DIM('a \ 'b) = card (Basis :: ('a \ 'b) set)" - unfolding dimension_prod_def Basis_prod_def - by (simp add: card_Un_disjoint disjoint_iff_not_equal - card_image inj_on_def dimension_def) -next - show "basis ` {.. 'b)} = (Basis :: ('a \ 'b) set)" - by (auto simp add: Basis_prod_def dimension_prod_def basis_prod_def - image_def elim!: Basis_elim) -next - show "basis ` {DIM('a \ 'b)..} = {0::('a \ 'b)}" - by (auto simp add: dimension_prod_def basis_prod_def prod_eq_iff image_def) qed +lemma DIM_prod[simp]: "DIM('a \ 'b) = DIM('a) + DIM('b)" + unfolding Basis_prod_def + by (subst card_Un_disjoint) (auto intro!: card_image arg_cong2[where f="op +"] inj_onI) + end end diff -r 46be26e02456 -r 899c9c4e4a4c src/HOL/Multivariate_Analysis/Fashoda.thy --- a/src/HOL/Multivariate_Analysis/Fashoda.thy Fri Dec 14 14:46:01 2012 +0100 +++ b/src/HOL/Multivariate_Analysis/Fashoda.thy Fri Dec 14 15:46:01 2012 +0100 @@ -7,6 +7,14 @@ imports Brouwer_Fixpoint Path_Connected Cartesian_Euclidean_Space begin +(* move *) + +lemma cart_eq_inner_axis: "a $ i = a \ axis i 1" + by (simp add: inner_axis) + +lemma axis_in_Basis: "a \ Basis \ axis i a \ Basis" + by (auto simp add: Basis_vec_def axis_eq_axis) + subsection {*Fashoda meet theorem. *} lemma infnorm_2: "infnorm (x::real^2) = max (abs(x$1)) (abs(x$2))" @@ -30,7 +38,7 @@ have lem1:"\z::real^2. infnorm(negatex z) = infnorm z" unfolding negatex_def infnorm_2 vector_2 by auto have lem2:"\z. z\0 \ infnorm(sqprojection z) = 1" unfolding sqprojection_def - unfolding infnorm_mul[unfolded smult_conv_scaleR] unfolding abs_inverse real_abs_infnorm + unfolding infnorm_mul[unfolded scalar_mult_eq_scaleR] unfolding abs_inverse real_abs_infnorm apply(subst infnorm_eq_0[THEN sym]) by auto let ?F = "(\w::real^2. (f \ (\x. x$1)) w - (g \ (\x. x$2)) w)" have *:"\i. (\x::real^2. x $ i) ` {- 1..1} = {- 1..1::real}" @@ -133,12 +141,6 @@ apply(rule_tac x="iscale s" in bexI) prefer 3 apply(rule_tac x="iscale t" in bexI) using isc[unfolded subset_eq, rule_format] by auto qed -(* move *) -lemma interval_bij_bij_cart: fixes x::"real^'n" assumes "\i. a$i < b$i \ u$i < v$i" - shows "interval_bij (a,b) (u,v) (interval_bij (u,v) (a,b) x) = x" - unfolding interval_bij_cart split_conv vec_eq_iff vec_lambda_beta - apply(rule,insert assms,erule_tac x=i in allE) by auto - lemma fashoda: fixes b::"real^2" assumes "path f" "path g" "path_image f \ {a..b}" "path_image g \ {a..b}" "(pathstart f)$1 = a$1" "(pathfinish f)$1 = b$1" @@ -184,8 +186,10 @@ "(interval_bij (a, b) (- 1, 1) \ f) 1 $ 1 = 1" "(interval_bij (a, b) (- 1, 1) \ g) 0 $ 2 = -1" "(interval_bij (a, b) (- 1, 1) \ g) 1 $ 2 = 1" - unfolding interval_bij_cart vector_component_simps o_def split_conv - unfolding assms[unfolded pathstart_def pathfinish_def] using as by auto qed note z=this + using assms as + by (simp_all add: axis_in_Basis cart_eq_inner_axis pathstart_def pathfinish_def interval_bij_def) + (simp_all add: inner_axis) + qed note z=this from z(1) guess zf unfolding image_iff .. note zf=this from z(2) guess zg unfolding image_iff .. note zg=this have *:"\i. (- 1) $ i < (1::real^2) $ i \ a $ i < b $ i" unfolding forall_2 using as by auto @@ -201,7 +205,7 @@ proof- let ?L = "\u. (x $ 1 = (1 - u) * a $ 1 + u * b $ 1 \ x $ 2 = (1 - u) * a $ 2 + u * b $ 2) \ 0 \ u \ u \ 1" { presume "?L \ ?R" "?R \ ?L" thus ?thesis unfolding closed_segment_def mem_Collect_eq - unfolding vec_eq_iff forall_2 smult_conv_scaleR[THEN sym] vector_component_simps by blast } + unfolding vec_eq_iff forall_2 scalar_mult_eq_scaleR[THEN sym] vector_component_simps by blast } { assume ?L then guess u apply-apply(erule exE)apply(erule conjE)+ . note u=this { fix b a assume "b + u * a > a + u * b" hence "(1 - u) * b > (1 - u) * a" by(auto simp add:field_simps) @@ -225,7 +229,7 @@ proof- let ?L = "\u. (x $ 1 = (1 - u) * a $ 1 + u * b $ 1 \ x $ 2 = (1 - u) * a $ 2 + u * b $ 2) \ 0 \ u \ u \ 1" { presume "?L \ ?R" "?R \ ?L" thus ?thesis unfolding closed_segment_def mem_Collect_eq - unfolding vec_eq_iff forall_2 smult_conv_scaleR[THEN sym] vector_component_simps by blast } + unfolding vec_eq_iff forall_2 scalar_mult_eq_scaleR[THEN sym] vector_component_simps by blast } { assume ?L then guess u apply-apply(erule exE)apply(erule conjE)+ . note u=this { fix b a assume "b + u * a > a + u * b" hence "(1 - u) * b > (1 - u) * a" by(auto simp add:field_simps) diff -r 46be26e02456 -r 899c9c4e4a4c src/HOL/Multivariate_Analysis/Finite_Cartesian_Product.thy --- a/src/HOL/Multivariate_Analysis/Finite_Cartesian_Product.thy Fri Dec 14 14:46:01 2012 +0100 +++ b/src/HOL/Multivariate_Analysis/Finite_Cartesian_Product.thy Fri Dec 14 15:46:01 2012 +0100 @@ -504,96 +504,11 @@ apply (simp add: axis_def) done -text {* A bijection between @{text "'n::finite"} and @{text "{.. ('n::finite)" where - "vec_bij_nat = (SOME p. bij_betw p {.. \ vec_bij_nat" -definition "\' = inv_into {..::nat \ ('n::finite))" - -lemma bij_betw_pi: - "bij_betw \ {..x. bij_betw x {..' (UNIV::'n set) {..'_def by auto - -lemma pi'_inj[intro]: "inj \'" - using bij_betw_pi' unfolding bij_betw_def by auto - -lemma pi'_range[intro]: "\i::'n. \' i < CARD('n::finite)" - using bij_betw_pi' unfolding bij_betw_def by auto - -lemma pi_pi'[simp]: "\i::'n::finite. \ (\' i) = i" - using bij_betw_pi by (auto intro!: f_inv_into_f simp: \'_def bij_betw_def) - -lemma pi'_pi[simp]: "\i. i\{.. \' (\ i::'n) = i" - using bij_betw_pi by (auto intro!: inv_into_f_eq simp: \'_def bij_betw_def) - -lemma pi_pi'_alt[simp]: "\i. i \' (\ i::'n) = i" - by auto - -lemma pi_inj_on: "inj_on (\::nat\'n::finite) {..(i div DIM('a))) (basis (i mod DIM('a))) - else 0)" - -lemma basis_eq: - assumes "i < CARD('b)" and "j < DIM('a)" - shows "basis (j + i * DIM('a)) = axis (\ i) (basis j)" -proof - - have "j + i * DIM('a) < DIM('a) * (i + 1)" using assms by (auto simp: field_simps) - also have "\ \ DIM('a) * CARD('b)" using assms unfolding mult_le_cancel1 by auto - finally show ?thesis - unfolding basis_vec_def using assms by (auto simp: vec_eq_iff not_less field_simps) -qed - -lemma basis_eq_pi': - assumes "j < DIM('a)" - shows "basis (j + \' i * DIM('a)) $ k = (if k = i then basis j else 0)" - apply (subst basis_eq) - using pi'_range assms by (simp_all add: axis_def) - -lemma split_times_into_modulo[consumes 1]: - fixes k :: nat - assumes "k < A * B" - obtains i j where "i < A" and "j < B" and "k = j + i * B" -proof - have "A * B \ 0" - proof assume "A * B = 0" with assms show False by simp qed - hence "0 < B" by auto - thus "k mod B < B" using `0 < B` by auto -next - have "k div B * B \ k div B * B + k mod B" by (rule le_add1) - also have "... < A * B" using assms by simp - finally show "k div B < A" by auto -qed simp - -lemma linear_less_than_times: - fixes i j A B :: nat assumes "i < B" "j < A" - shows "j + i * A < B * A" -proof - - have "i * A + j < (Suc i)*A" using `j < A` by simp - also have "\ \ B * A" using `i < B` unfolding mult_le_cancel2 by simp - finally show ?thesis by simp -qed - -lemma DIM_cart[simp]: "DIM('a^'b) = CARD('b) * DIM('a)" - by (rule dimension_vec_def) - instance proof show "(Basis :: ('a ^ 'b) set) \ {}" unfolding Basis_vec_def by simp @@ -611,27 +526,17 @@ show "(\u\Basis. inner x u = 0) \ x = 0" unfolding Basis_vec_def by (simp add: inner_axis euclidean_all_zero_iff vec_eq_iff) -next - show "DIM('a ^ 'b) = card (Basis :: ('a ^ 'b) set)" - unfolding Basis_vec_def dimension_vec_def dimension_def - by (simp add: card_UN_disjoint [unfolded disjoint_iff_not_equal] - axis_eq_axis nonzero_Basis) -next - show "basis ` {..x::'a::euclidean_space. x $$ k)" - apply (rule bounded_linearI[where K=1]) - using component_le_norm[of _ k] - unfolding real_norm_def - apply auto - done +lemma bounded_linear_component [intro]: "bounded_linear (\x::'a::euclidean_space. x \ k)" + by (rule bounded_linear_inner_left) lemma transitive_stepwise_lt_eq: assumes "(\x y z::nat. R x y \ R y z \ R x z)" @@ -162,16 +158,10 @@ subsection {* Some useful lemmas about intervals. *} -abbreviation One where "One \ ((\\ i. 1)::_::ordered_euclidean_space)" - -lemma empty_as_interval: "{} = {One..0}" - apply (rule set_eqI, rule) - defer - unfolding mem_interval - using UNIV_witness[where 'a='n] - apply (erule_tac exE, rule_tac x = x in allE) - apply auto - done +abbreviation One where "One \ ((\Basis)::_::euclidean_space)" + +lemma empty_as_interval: "{} = {One..(0::'a::ordered_euclidean_space)}" + by (auto simp: set_eq_iff eucl_le[where 'a='a] intro!: bexI[OF _ SOME_Basis]) lemma interior_subset_union_intervals: assumes "i = {a..b::'a::ordered_euclidean_space}" "j = {c..d}" @@ -255,17 +245,17 @@ done next case False - then obtain k where "x$$k \ a$$k \ x$$k \ b$$k" and k:"kk \ a\k \ x\k \ b\k" and k:"k\Basis" unfolding mem_interval by (auto simp add: not_less) - hence "x$$k = a$$k \ x$$k = b$$k" + hence "x\k = a\k \ x\k = b\k" using True unfolding ab and mem_interval - apply (erule_tac x = k in allE) + apply (erule_tac x = k in ballE) apply auto done hence "\x. ball x (e/2) \ s \ (\f)" proof (erule_tac disjE) - let ?z = "x - (e/2) *\<^sub>R basis k" - assume as: "x$$k = a$$k" + let ?z = "x - (e/2) *\<^sub>R k" + assume as: "x\k = a\k" have "ball ?z (e / 2) \ i = {}" apply (rule ccontr) unfolding ex_in_conv[THEN sym] @@ -273,15 +263,12 @@ fix y assume "y \ ball ?z (e / 2) \ i" hence "dist ?z y < e/2" and yi:"y\i" by auto - hence "\(?z - y) $$ k\ < e/2" - using component_le_norm[of "?z - y" k] unfolding dist_norm by auto - hence "y$$k < a$$k" - using e[THEN conjunct1] k by (auto simp add: field_simps as) + hence "\(?z - y) \ k\ < e/2" + using Basis_le_norm[OF k, of "?z - y"] unfolding dist_norm by auto + hence "y\k < a\k" + using e[THEN conjunct1] k by (auto simp add: field_simps as inner_Basis inner_simps) hence "y \ i" - unfolding ab mem_interval not_all - apply (rule_tac x=k in exI) - using k apply auto - done + unfolding ab mem_interval by (auto intro!: bexI[OF _ k]) thus False using yi by auto qed moreover @@ -290,10 +277,10 @@ proof fix y assume as: "y\ ball ?z (e/2)" - have "norm (x - y) \ \e\ / 2 + norm (x - y - (e / 2) *\<^sub>R basis k)" + have "norm (x - y) \ \e\ / 2 + norm (x - y - (e / 2) *\<^sub>R k)" apply - - apply (rule order_trans,rule norm_triangle_sub[of "x - y" "(e/2) *\<^sub>R basis k"]) - unfolding norm_scaleR norm_basis + apply (rule order_trans,rule norm_triangle_sub[of "x - y" "(e/2) *\<^sub>R k"]) + unfolding norm_scaleR norm_Basis[OF k] apply auto done also have "\ < \e\ / 2 + \e\ / 2" @@ -310,8 +297,8 @@ apply auto done next - let ?z = "x + (e/2) *\<^sub>R basis k" - assume as: "x$$k = b$$k" + let ?z = "x + (e/2) *\<^sub>R k" + assume as: "x\k = b\k" have "ball ?z (e / 2) \ i = {}" apply (rule ccontr) unfolding ex_in_conv[THEN sym] @@ -319,15 +306,12 @@ fix y assume "y \ ball ?z (e / 2) \ i" hence "dist ?z y < e/2" and yi:"y\i" by auto - hence "\(?z - y) $$ k\ < e/2" - using component_le_norm[of "?z - y" k] unfolding dist_norm by auto - hence "y$$k > b$$k" - using e[THEN conjunct1] k by(auto simp add:field_simps as) + hence "\(?z - y) \ k\ < e/2" + using Basis_le_norm[OF k, of "?z - y"] unfolding dist_norm by auto + hence "y\k > b\k" + using e[THEN conjunct1] k by(auto simp add:field_simps inner_simps inner_Basis as) hence "y \ i" - unfolding ab mem_interval not_all - using k apply (rule_tac x=k in exI) - apply auto - done + unfolding ab mem_interval by (auto intro!: bexI[OF _ k]) thus False using yi by auto qed moreover @@ -336,11 +320,11 @@ proof fix y assume as: "y\ ball ?z (e/2)" - have "norm (x - y) \ \e\ / 2 + norm (x - y + (e / 2) *\<^sub>R basis k)" + have "norm (x - y) \ \e\ / 2 + norm (x - y + (e / 2) *\<^sub>R k)" apply - - apply(rule order_trans,rule norm_triangle_sub[of "x - y" "- (e/2) *\<^sub>R basis k"]) + apply(rule order_trans,rule norm_triangle_sub[of "x - y" "- (e/2) *\<^sub>R k"]) unfolding norm_scaleR - apply auto + apply (auto simp: k) done also have "\ < \e\ / 2 + \e\ / 2" apply (rule add_strict_left_mono) @@ -383,61 +367,25 @@ subsection {* Bounds on intervals where they exist. *} -definition "interval_upperbound (s::('a::ordered_euclidean_space) set) = - ((\\ i. Sup {a. \x\s. x$$i = a})::'a)" - -definition "interval_lowerbound (s::('a::ordered_euclidean_space) set) = - ((\\ i. Inf {a. \x\s. x$$i = a})::'a)" +definition interval_upperbound :: "('a::ordered_euclidean_space) set \ 'a" where + "interval_upperbound s = (\i\Basis. Sup {a. \x\s. x\i = a} *\<^sub>R i)" + +definition interval_lowerbound :: "('a::ordered_euclidean_space) set \ 'a" where + "interval_lowerbound s = (\i\Basis. Inf {a. \x\s. x\i = a} *\<^sub>R i)" lemma interval_upperbound[simp]: - assumes "\i (b::'a)$$i" - shows "interval_upperbound {a..b} = b" - using assms - unfolding interval_upperbound_def - apply (subst euclidean_eq[where 'a='a]) - apply safe - unfolding euclidean_lambda_beta' - apply (erule_tac x=i in allE) - apply (rule Sup_unique) - unfolding setle_def - apply rule - unfolding mem_Collect_eq - apply (erule bexE) - unfolding mem_interval - defer - apply (rule, rule) - apply (rule_tac x="b$$i" in bexI) - defer - unfolding mem_Collect_eq - apply (rule_tac x=b in bexI) - unfolding mem_interval - using assms apply auto - done + "\i\Basis. a\i \ b\i \ + interval_upperbound {a..b} = (b::'a::ordered_euclidean_space)" + unfolding interval_upperbound_def euclidean_representation_setsum + by (auto simp del: ex_simps simp add: Bex_def ex_simps[symmetric] eucl_le[where 'a='a] setle_def + intro!: Sup_unique) lemma interval_lowerbound[simp]: - assumes "\i (b::'a)$$i" - shows "interval_lowerbound {a..b} = a" - using assms - unfolding interval_lowerbound_def - apply (subst euclidean_eq[where 'a='a]) - apply safe - unfolding euclidean_lambda_beta' - apply (erule_tac x=i in allE) - apply (rule Inf_unique) - unfolding setge_def - apply rule - unfolding mem_Collect_eq - apply (erule bexE) - unfolding mem_interval - defer - apply (rule, rule) - apply (rule_tac x = "a$$i" in bexI) - defer - unfolding mem_Collect_eq - apply (rule_tac x=a in bexI) - unfolding mem_interval - using assms apply auto - done + "\i\Basis. a\i \ b\i \ + interval_lowerbound {a..b} = (a::'a::ordered_euclidean_space)" + unfolding interval_lowerbound_def euclidean_representation_setsum + by (auto simp del: ex_simps simp add: Bex_def ex_simps[symmetric] eucl_le[where 'a='a] setge_def + intro!: Inf_unique) lemmas interval_bounds = interval_upperbound interval_lowerbound @@ -449,15 +397,15 @@ subsection {* Content (length, area, volume...) of an interval. *} definition "content (s::('a::ordered_euclidean_space) set) = - (if s = {} then 0 else (\ii b$$i \ {a..b::'a::ordered_euclidean_space} \ {}" + (if s = {} then 0 else (\i\Basis. (interval_upperbound s)\i - (interval_lowerbound s)\i))" + +lemma interval_not_empty:"\i\Basis. a\i \ b\i \ {a..b::'a::ordered_euclidean_space} \ {}" unfolding interval_eq_empty unfolding not_ex not_less by auto lemma content_closed_interval: fixes a::"'a::ordered_euclidean_space" - assumes "\i b$$i" - shows "content {a..b} = (\ii\Basis. a\i \ b\i" + shows "content {a..b} = (\i\Basis. b\i - a\i)" using interval_not_empty[OF assms] unfolding content_def interval_upperbound[OF assms] interval_lowerbound[OF assms] by auto @@ -465,7 +413,7 @@ lemma content_closed_interval': fixes a::"'a::ordered_euclidean_space" assumes "{a..b}\{}" - shows "content {a..b} = (\ii\Basis. b\i - a\i)" apply (rule content_closed_interval) using assms unfolding interval_ne_empty apply assumption @@ -482,13 +430,13 @@ lemma content_singleton[simp]: "content {a} = 0" proof - have "content {a .. a} = 0" - by (subst content_closed_interval) auto + by (subst content_closed_interval) (auto simp: ex_in_conv) then show ?thesis by simp qed lemma content_unit[intro]: "content{0..One::'a::ordered_euclidean_space} = 1" proof - - have *: "\i (One::'a)$$i" by auto + have *: "\i\Basis. (0::'a)\i \ (One::'a)\i" by auto have "0 \ {0..One::'a}" unfolding mem_interval by auto thus ?thesis unfolding content_def interval_bounds[OF *] using setprod_1 by auto qed @@ -498,12 +446,12 @@ shows "0 \ content {a..b}" proof (cases "{a..b} = {}") case False - hence *: "\i b $$ i" unfolding interval_ne_empty . - have "(\i 0" + hence *: "\i\Basis. a \ i \ b \ i" unfolding interval_ne_empty . + have "(\i\Basis. interval_upperbound {a..b} \ i - interval_lowerbound {a..b} \ i) \ 0" apply (rule setprod_nonneg) unfolding interval_bounds[OF *] using * - apply (erule_tac x=x in allE) + apply (erule_tac x=x in ballE) apply auto done thus ?thesis unfolding content_def by (auto simp del:interval_bounds') @@ -511,75 +459,59 @@ lemma content_pos_lt: fixes a::"'a::ordered_euclidean_space" - assumes "\ii\Basis. a\i < b\i" shows "0 < content {a..b}" proof - - have help_lemma1: "\i \i ((b$$i)::real)" - apply (rule, erule_tac x=i in allE) + have help_lemma1: "\i\Basis. a\i < b\i \ \i\Basis. a\i \ ((b\i)::real)" + apply (rule, erule_tac x=i in ballE) apply auto done show ?thesis unfolding content_closed_interval[OF help_lemma1[OF assms]] apply(rule setprod_pos) - using assms apply (erule_tac x=x in allE) + using assms apply (erule_tac x=x in ballE) apply auto done qed -lemma content_eq_0: "content{a..b::'a::ordered_euclidean_space} = 0 \ (\i a$$i)" +lemma content_eq_0: "content{a..b::'a::ordered_euclidean_space} = 0 \ (\i\Basis. b\i \ a\i)" proof (cases "{a..b} = {}") case True thus ?thesis unfolding content_def if_P[OF True] unfolding interval_eq_empty apply - - apply (rule, erule exE) - apply (rule_tac x = i in exI) + apply (rule, erule bexE) + apply (rule_tac x = i in bexI) apply auto done next case False from this[unfolded interval_eq_empty not_ex not_less] - have as: "\i a $$ i" by fastforce + have as: "\i\Basis. b \ i \ a \ i" by fastforce show ?thesis - unfolding content_def if_not_P[OF False] setprod_zero_iff[OF finite_lessThan] - apply rule - apply (erule_tac[!] exE bexE) - unfolding interval_bounds[OF as] - apply (rule_tac x=x in exI) - defer - apply (rule_tac x=i in bexI) - using as apply (erule_tac x=i in allE) - apply auto - done + unfolding content_def if_not_P[OF False] setprod_zero_iff[OF finite_Basis] + using as + by (auto intro!: bexI) qed lemma cond_cases:"(P \ Q x) \ (\ P \ Q y) \ Q (if P then x else y)" by auto lemma content_closed_interval_cases: "content {a..b::'a::ordered_euclidean_space} = - (if \i b$$i then setprod (\i. b$$i - a$$i) {..i\Basis. a\i \ b\i then setprod (\i. b\i - a\i) Basis else 0)" + by (auto simp: not_le content_eq_0 intro: less_imp_le content_closed_interval) lemma content_eq_0_interior: "content {a..b} = 0 \ interior({a..b}) = {}" unfolding content_eq_0 interior_closed_interval interval_eq_empty by auto -(*lemma content_eq_0_1: "content {a..b::real^1} = 0 \ dest_vec1 b \ dest_vec1 a" - unfolding content_eq_0 by auto*) - -lemma content_pos_lt_eq: "0 < content {a..b::'a::ordered_euclidean_space} \ (\i (\i\Basis. a\i < b\i)" apply rule defer apply (rule content_pos_lt, assumption) proof - assume "0 < content {a..b}" hence "content {a..b} \ 0" by auto - thus "\ii\Basis. a\i < b\i" unfolding content_eq_0 not_ex not_le by fastforce qed @@ -593,20 +525,20 @@ thus ?thesis using content_pos_le[of c d] by auto next case False - hence ab_ne:"\i b $$ i" unfolding interval_ne_empty by auto + hence ab_ne:"\i\Basis. a \ i \ b \ i" unfolding interval_ne_empty by auto hence ab_ab:"a\{a..b}" "b\{a..b}" unfolding mem_interval by auto have "{c..d} \ {}" using assms False by auto - hence cd_ne:"\i d $$ i" using assms unfolding interval_ne_empty by auto + hence cd_ne:"\i\Basis. c \ i \ d \ i" using assms unfolding interval_ne_empty by auto show ?thesis unfolding content_def unfolding interval_bounds[OF ab_ne] interval_bounds[OF cd_ne] unfolding if_not_P[OF False] if_not_P[OF `{c..d} \ {}`] apply(rule setprod_mono,rule) proof - fix i - assume i:"i\{.. b $$ i - a $$ i" using ab_ne[THEN spec[where x=i]] i by auto - show "b $$ i - a $$ i \ d $$ i - c $$ i" + fix i :: 'a + assume i:"i\Basis" + show "0 \ b \ i - a \ i" using ab_ne[THEN bspec, OF i] i by auto + show "b \ i - a \ i \ d \ i - c \ i" using assms[unfolded subset_eq mem_interval,rule_format,OF ab_ab(2),of i] using assms[unfolded subset_eq mem_interval,rule_format,OF ab_ab(1),of i] using i by auto @@ -857,134 +789,84 @@ show "k \ {}" using k d1(3) d2(3) by auto show "\a b. k = {a..b}" using k d1(4) d2(4) by auto qed lemma partial_division_extend_1: - assumes "{c..d} \ {a..b::'a::ordered_euclidean_space}" "{c..d} \ {}" + assumes incl: "{c..d} \ {a..b::'a::ordered_euclidean_space}" and nonempty: "{c..d} \ {}" obtains p where "p division_of {a..b}" "{c..d} \ p" -proof- def n \ "DIM('a)" have n:"1 \ n" "0 < n" "n \ 0" unfolding n_def using DIM_positive[where 'a='a] by auto - guess \ using ex_bij_betw_nat_finite_1[OF finite_lessThan[of "DIM('a)"]] .. note \=this - def \' \ "inv_into {1..n} \" - have \':"bij_betw \' {..] unfolding \'_def n_def by auto - hence \'_i:"\i. i \' i \ {1..n}" unfolding bij_betw_def by auto - have \_i:"\i. i\{1..n} \ \ i unfolding bij_betw_def n_def by auto - have \_\'[simp]:"\i. i \ (\' i) = i" unfolding \'_def - apply(rule f_inv_into_f) unfolding n_def using \ unfolding bij_betw_def by auto - have \'_\[simp]:"\i. i\{1..n} \ \' (\ i) = i" unfolding \'_def apply(rule inv_into_f_eq) - using \ unfolding n_def bij_betw_def by auto - have "{c..d} \ {}" using assms by auto - let ?p1 = "\l. {(\\ i. if \' i < l then c$$i else a$$i)::'a .. (\\ i. if \' i < l then d$$i else if \' i = l then c$$\ l else b$$i)}" - let ?p2 = "\l. {(\\ i. if \' i < l then c$$i else if \' i = l then d$$\ l else a$$i)::'a .. (\\ i. if \' i < l then d$$i else b$$i)}" - let ?p = "{?p1 l |l. l \ {1..n+1}} \ {?p2 l |l. l \ {1..n+1}}" - have abcd:"\i. i a $$ i \ c $$ i \ c$$i \ d$$i \ d $$ i \ b $$ i" using assms - unfolding subset_interval interval_eq_empty by auto - show ?thesis apply(rule that[of ?p]) apply(rule division_ofI) - proof- have "\i. i \' i < Suc n" - proof(rule ccontr,unfold not_less) fix i assume i:"i \' i" - hence "\' i \ {1..n}" by auto thus False using \' i unfolding bij_betw_def by auto - qed hence "c = (\\ i. if \' i < Suc n then c $$ i else a $$ i)" - "d = (\\ i. if \' i < Suc n then d $$ i else if \' i = n + 1 then c $$ \ (n + 1) else b $$ i)" - unfolding euclidean_eq[where 'a='a] using \' unfolding bij_betw_def by auto - thus cdp:"{c..d} \ ?p" apply-apply(rule UnI1) unfolding mem_Collect_eq apply(rule_tac x="n + 1" in exI) by auto - have "\l. l\{1..n+1} \ ?p1 l \ {a..b}" "\l. l\{1..n+1} \ ?p2 l \ {a..b}" - unfolding subset_eq apply(rule_tac[!] ballI,rule_tac[!] ccontr) - proof- fix l assume l:"l\{1..n+1}" fix x assume "x\{a..b}" - then guess i unfolding mem_interval not_all not_imp .. note i=conjunctD2[OF this] - show "x \ ?p1 l \ False" "x \ ?p2 l \ False" unfolding mem_interval apply(erule_tac[!] x=i in allE) - apply(case_tac[!] "\' i < l", case_tac[!] "\' i = l") using abcd[of i] i by auto - qed moreover have "\x. x \ {a..b} \ x \ \?p" - proof- fix x assume x:"x\{a..b}" - { presume "x\{c..d} \ x \ \?p" thus "x \ \?p" using cdp by blast } - let ?M = "{i. i\{1..n+1} \ \ (c $$ \ i \ x $$ \ i \ x $$ \ i \ d $$ \ i)}" - assume "x\{c..d}" then guess i0 unfolding mem_interval not_all not_imp .. - hence "\' i0 \ ?M" using \' unfolding bij_betw_def by(auto intro!:le_SucI) - hence M:"finite ?M" "?M \ {}" by auto - def l \ "Min ?M" note l = Min_less_iff[OF M,unfolded l_def[symmetric]] Min_in[OF M,unfolded mem_Collect_eq l_def[symmetric]] - Min_gr_iff[OF M,unfolded l_def[symmetric]] - have "x\?p1 l \ x\?p2 l" using l(2)[THEN conjunct2] unfolding de_Morgan_conj not_le - apply- apply(erule disjE) apply(rule disjI1) defer apply(rule disjI2) - proof- assume as:"x $$ \ l < c $$ \ l" - show "x \ ?p1 l" unfolding mem_interval apply safe unfolding euclidean_lambda_beta' - proof- case goal1 have "\' i \ {1..n}" using \' unfolding bij_betw_def not_le using goal1 by auto - thus ?case using as x[unfolded mem_interval,rule_format,of i] - apply auto using l(3)[of "\' i"] using goal1 by(auto elim!:ballE[where x="\' i"]) - next case goal2 have "\' i \ {1..n}" using \' unfolding bij_betw_def not_le using goal2 by auto - thus ?case using as x[unfolded mem_interval,rule_format,of i] - apply auto using l(3)[of "\' i"] using goal2 by(auto elim!:ballE[where x="\' i"]) - qed - next assume as:"x $$ \ l > d $$ \ l" - show "x \ ?p2 l" unfolding mem_interval apply safe unfolding euclidean_lambda_beta' - proof- fix i assume i:"i' i \ {1..n}" using \' unfolding bij_betw_def not_le using i by auto - thus "(if \' i < l then c $$ i else if \' i = l then d $$ \ l else a $$ i) \ x $$ i" - "x $$ i \ (if \' i < l then d $$ i else b $$ i)" - using as x[unfolded mem_interval,rule_format,of i] - apply auto using l(3)[of "\' i"] i by(auto elim!:ballE[where x="\' i"]) - qed qed - thus "x \ \?p" using l(2) by blast - qed ultimately show "\?p = {a..b}" apply-apply(rule) defer apply(rule) by(assumption,blast) - - show "finite ?p" by auto - fix k assume k:"k\?p" then obtain l where l:"k = ?p1 l \ k = ?p2 l" "l \ {1..n + 1}" by auto - show "k\{a..b}" apply(rule,unfold mem_interval,rule,rule) - proof fix i x assume i:"i k" moreover have "\' i < l \ \' i = l \ \' i > l" by auto - ultimately show "a$$i \ x$$i" "x$$i \ b$$i" using abcd[of i] using l using i - by(auto elim!:allE[where x=i] simp add:eucl_le[where 'a='a]) (* FIXME: SLOW *) - qed have "\l. ?p1 l \ {}" "\l. ?p2 l \ {}" unfolding interval_eq_empty not_ex apply(rule_tac[!] allI) - proof- case goal1 thus ?case using abcd[of x] by auto - next case goal2 thus ?case using abcd[of x] by auto - qed thus "k \ {}" using k by auto - show "\a b. k = {a..b}" using k by auto - fix k' assume k':"k' \ ?p" "k \ k'" then obtain l' where l':"k' = ?p1 l' \ k' = ?p2 l'" "l' \ {1..n + 1}" by auto - { fix k k' l l' - assume k:"k\?p" and l:"k = ?p1 l \ k = ?p2 l" "l \ {1..n + 1}" - assume k':"k' \ ?p" "k \ k'" and l':"k' = ?p1 l' \ k' = ?p2 l'" "l' \ {1..n + 1}" - assume "l \ l'" fix x - have "x \ interior k \ interior k'" - proof(rule,cases "l' = n+1") assume x:"x \ interior k \ interior k'" - case True hence "\i. i \' i < l'" using \'_i using l' by(auto simp add:less_Suc_eq_le) - hence *:"\ P Q. (\\ i. if \' i < l' then P i else Q i) = ((\\ i. P i)::'a)" apply-apply(subst euclidean_eq) by auto - hence k':"k' = {c..d}" using l'(1) unfolding * by auto - have ln:"l < n + 1" - proof(rule ccontr) case goal1 hence l2:"l = n+1" using l by auto - hence "\i. i \' i < l" using \'_i by(auto simp add:less_Suc_eq_le) - hence *:"\ P Q. (\\ i. if \' i < l then P i else Q i) = ((\\ i. P i)::'a)" apply-apply(subst euclidean_eq) by auto - hence "k = {c..d}" using l(1) \'_i unfolding * by(auto) - thus False using `k\k'` k' by auto - qed have **:"\' (\ l) = l" using \'_\[of l] using l ln by auto - have "x $$ \ l < c $$ \ l \ d $$ \ l < x $$ \ l" using l(1) apply- - proof(erule disjE) - assume as:"k = ?p1 l" note * = conjunct1[OF x[unfolded as Int_iff interior_closed_interval mem_interval],rule_format] - show ?thesis using *[of "\ l"] using ln l(2) using \_i[of l] by(auto simp add:** not_less) - next assume as:"k = ?p2 l" note * = conjunct1[OF x[unfolded as Int_iff interior_closed_interval mem_interval],rule_format] - show ?thesis using *[of "\ l"] using ln l(2) using \_i[of l] unfolding ** by auto - qed thus False using x unfolding k' unfolding Int_iff interior_closed_interval mem_interval - by(auto elim!:allE[where x="\ l"]) - next case False hence "l < n + 1" using l'(2) using `l\l'` by auto - hence ln:"l \ {1..n}" "l' \ {1..n}" using l l' False by auto - note \_l = \'_\[OF ln(1)] \'_\[OF ln(2)] - assume x:"x \ interior k \ interior k'" - show False using l(1) l'(1) apply- - proof(erule_tac[!] disjE)+ - assume as:"k = ?p1 l" "k' = ?p1 l'" - note * = conjunctD2[OF x[unfolded as Int_iff interior_closed_interval mem_interval],rule_format] - have "l \ l'" using k'(2)[unfolded as] by auto - thus False using *[of "\ l'"] *[of "\ l"] ln using \_i[OF ln(1)] \_i[OF ln(2)] apply(cases "l_l \_i n_def) - next assume as:"k = ?p2 l" "k' = ?p2 l'" - note * = conjunctD2[OF x[unfolded as Int_iff interior_closed_interval mem_interval],rule_format] - have "l \ l'" apply(rule) using k'(2)[unfolded as] by auto - thus False using *[of "\ l"] *[of "\ l'"] `l \ l'` ln by(auto simp add:euclidean_lambda_beta' \_l \_i n_def) - next assume as:"k = ?p1 l" "k' = ?p2 l'" - note * = conjunctD2[OF x[unfolded as Int_iff interior_closed_interval mem_interval],rule_format] - show False using abcd[of "\ l'"] using *[of "\ l"] *[of "\ l'"] `l \ l'` ln apply(cases "l=l'") - by(auto simp add:euclidean_lambda_beta' \_l \_i n_def) - next assume as:"k = ?p2 l" "k' = ?p1 l'" - note * = conjunctD2[OF x[unfolded as Int_iff interior_closed_interval mem_interval],rule_format] - show False using *[of "\ l"] *[of "\ l'"] ln `l \ l'` apply(cases "l=l'") using abcd[of "\ l'"] - by(auto simp add:euclidean_lambda_beta' \_l \_i n_def) - qed qed } - from this[OF k l k' l'] this[OF k'(1) l' k _ l] have "\x. x \ interior k \ interior k'" - apply - apply(cases "l' \ l") using k'(2) by auto - thus "interior k \ interior k' = {}" by auto -qed qed +proof + let ?B = "\f::'a\'a \ 'a. {(\i\Basis. (fst (f i) \ i) *\<^sub>R i) .. (\i\Basis. (snd (f i) \ i) *\<^sub>R i)}" + def p \ "?B ` (Basis \\<^isub>E {(a, c), (c, d), (d, b)})" + + show "{c .. d} \ p" + unfolding p_def + by (auto simp add: interval_eq_empty eucl_le[where 'a='a] + intro!: image_eqI[where x="\(i::'a)\Basis. (c, d)"]) + + { fix i :: 'a assume "i \ Basis" + with incl nonempty have "a \ i \ c \ i" "c \ i \ d \ i" "d \ i \ b \ i" + unfolding interval_eq_empty subset_interval by (auto simp: not_le) } + note ord = this + + show "p division_of {a..b}" + proof (rule division_ofI) + show "finite p" + unfolding p_def by (auto intro!: finite_PiE) + { fix k assume "k \ p" + then obtain f where f: "f \ Basis \\<^isub>E {(a, c), (c, d), (d, b)}" and k: "k = ?B f" + by (auto simp: p_def) + then show "\a b. k = {a..b}" by auto + have "k \ {a..b} \ k \ {}" + proof (simp add: k interval_eq_empty subset_interval not_less, safe) + fix i :: 'a assume i: "i \ Basis" + moreover with f have "f i = (a, c) \ f i = (c, d) \ f i = (d, b)" + by (auto simp: PiE_iff) + moreover note ord[of i] + ultimately show "a \ i \ fst (f i) \ i" "snd (f i) \ i \ b \ i" "fst (f i) \ i \ snd (f i) \ i" + by (auto simp: subset_iff eucl_le[where 'a='a]) + qed + then show "k \ {}" "k \ {a .. b}" by auto + { + fix l assume "l \ p" + then obtain g where g: "g \ Basis \\<^isub>E {(a, c), (c, d), (d, b)}" and l: "l = ?B g" + by (auto simp: p_def) + assume "l \ k" + have "\i\Basis. f i \ g i" + proof (rule ccontr) + assume "\ (\i\Basis. f i \ g i)" + with f g have "f = g" + by (auto simp: PiE_iff extensional_def intro!: ext) + with `l \ k` show False + by (simp add: l k) + qed + then guess i .. + moreover then have "f i = (a, c) \ f i = (c, d) \ f i = (d, b)" + "g i = (a, c) \ g i = (c, d) \ g i = (d, b)" + using f g by (auto simp: PiE_iff) + moreover note ord[of i] + ultimately show "interior l \ interior k = {}" + by (auto simp add: l k interior_closed_interval disjoint_interval intro!: bexI[of _ i]) } + note `k \ { a.. b}` } + moreover + { fix x assume x: "x \ {a .. b}" + have "\i\Basis. \l. x \ i \ {fst l \ i .. snd l \ i} \ l \ {(a, c), (c, d), (d, b)}" + proof + fix i :: 'a assume "i \ Basis" + with x ord[of i] + have "(a \ i \ x \ i \ x \ i \ c \ i) \ (c \ i \ x \ i \ x \ i \ d \ i) \ + (d \ i \ x \ i \ x \ i \ b \ i)" + by (auto simp: eucl_le[where 'a='a]) + then show "\l. x \ i \ {fst l \ i .. snd l \ i} \ l \ {(a, c), (c, d), (d, b)}" + by auto + qed + then guess f unfolding bchoice_iff .. note f = this + moreover then have "restrict f Basis \ Basis \\<^isub>E {(a, c), (c, d), (d, b)}" + by auto + moreover from f have "x \ ?B (restrict f Basis)" + by (auto simp: mem_interval eucl_le[where 'a='a]) + ultimately have "\k\p. x \ k" + unfolding p_def by blast } + ultimately show "\p = {a..b}" + by auto + qed +qed lemma partial_division_extend_interval: assumes "p division_of (\p)" "(\p) \ {a..b}" obtains q where "p \ q" "q division_of {a..b::'a::ordered_euclidean_space}" proof(cases "p = {}") @@ -1415,9 +1297,9 @@ lemma interval_bisection_step: fixes type::"'a::ordered_euclidean_space" assumes "P {}" "(\s t. P s \ P t \ interior(s) \ interior(t) = {} \ P(s \ t))" "~(P {a..b::'a})" obtains c d where "~(P{c..d})" - "\i c$$i \ c$$i \ d$$i \ d$$i \ b$$i \ 2 * (d$$i - c$$i) \ b$$i - a$$i" + "\i\Basis. a\i \ c\i \ c\i \ d\i \ d\i \ b\i \ 2 * (d\i - c\i) \ b\i - a\i" proof- have "{a..b} \ {}" using assms(1,3) by auto - note ab=this[unfolded interval_eq_empty not_ex not_less] + then have ab: "\i. i\Basis \ a \ i \ b \ i" by (auto simp: interval_eq_empty not_le) { fix f have "finite f \ (\s\f. P s) \ (\s\f. \a b. s = {a..b}) \ @@ -1428,60 +1310,72 @@ apply rule defer apply rule defer apply(rule inter_interior_unions_intervals) using insert by auto qed } note * = this - let ?A = "{{c..d} | c d::'a. \i (d$$i = (a$$i + b$$i) / 2) \ (c$$i = (a$$i + b$$i) / 2) \ (d$$i = b$$i)}" - let ?PP = "\c d. \i c$$i \ c$$i \ d$$i \ d$$i \ b$$i \ 2 * (d$$i - c$$i) \ b$$i - a$$i" + let ?A = "{{c..d} | c d::'a. \i\Basis. (c\i = a\i) \ (d\i = (a\i + b\i) / 2) \ (c\i = (a\i + b\i) / 2) \ (d\i = b\i)}" + let ?PP = "\c d. \i\Basis. a\i \ c\i \ c\i \ d\i \ d\i \ b\i \ 2 * (d\i - c\i) \ b\i - a\i" { presume "\c d. ?PP c d \ P {c..d} \ False" thus thesis unfolding atomize_not not_all apply-apply(erule exE)+ apply(rule_tac c=x and d=xa in that) by auto } assume as:"\c d. ?PP c d \ P {c..d}" have "P (\ ?A)" proof(rule *, rule_tac[2-] ballI, rule_tac[4] ballI, rule_tac[4] impI) - let ?B = "(\s.{(\\ i. if i \ s then a$$i else (a$$i + b$$i) / 2)::'a .. - (\\ i. if i \ s then (a$$i + b$$i) / 2 else b$$i)}) ` {s. s \ {.. ?B" proof case goal1 then guess c unfolding mem_Collect_eq .. then guess d apply- by(erule exE,(erule conjE)+) note c_d=this[rule_format] have *:"\a b c d. a = c \ b = d \ {a..b} = {c..d}" by auto - show "x\?B" unfolding image_iff apply(rule_tac x="{i. i c$$i = a$$i}" in bexI) - unfolding c_d apply(rule * ) unfolding euclidean_eq[where 'a='a] apply safe unfolding euclidean_lambda_beta' mem_Collect_eq - proof- fix i assume "i c $$ i = a $$ i then a $$ i else (a $$ i + b $$ i) / 2)" - "d $$ i = (if i < DIM('a) \ c $$ i = a $$ i then (a $$ i + b $$ i) / 2 else b $$ i)" - using c_d(2)[of i] ab[THEN spec[where x=i]] by(auto simp add:field_simps) + show "x\?B" unfolding image_iff + apply(rule_tac x="{i. i\Basis \ c\i = a\i}" in bexI) + unfolding c_d + apply(rule *) + apply (simp_all only: euclidean_eq_iff[where 'a='a] inner_setsum_left_Basis mem_Collect_eq simp_thms + cong: ball_cong) + apply safe + proof- + fix i :: 'a assume i: "i\Basis" + thus " c \ i = (if c \ i = a \ i then a \ i else (a \ i + b \ i) / 2)" + "d \ i = (if c \ i = a \ i then (a \ i + b \ i) / 2 else b \ i)" + using c_d(2)[of i] ab[OF i] by(auto simp add:field_simps) qed qed thus "finite ?A" apply(rule finite_subset) by auto fix s assume "s\?A" then guess c unfolding mem_Collect_eq .. then guess d apply- by(erule exE,(erule conjE)+) note c_d=this[rule_format] show "P s" unfolding c_d apply(rule as[rule_format]) proof- case goal1 thus ?case - using c_d(2)[of i] using ab[THEN spec[where x=i]] by auto qed + using c_d(2)[of i] using ab[OF `i \ Basis`] by auto qed show "\a b. s = {a..b}" unfolding c_d by auto fix t assume "t\?A" then guess e unfolding mem_Collect_eq .. then guess f apply- by(erule exE,(erule conjE)+) note e_f=this[rule_format] assume "s \ t" hence "\ (c = e \ d = f)" unfolding c_d e_f by auto - then obtain i where "c$$i \ e$$i \ d$$i \ f$$i" and i':"iBasis" + unfolding euclidean_eq_iff[where 'a='a] by auto + hence i:"c\i \ e\i" "d\i \ f\i" apply- apply(erule_tac[!] disjE) + proof- assume "c\i \ e\i" thus "d\i \ f\i" using c_d(2)[OF i'] e_f(2)[OF i'] by fastforce + next assume "d\i \ f\i" thus "c\i \ e\i" using c_d(2)[OF i'] e_f(2)[OF i'] by fastforce qed have *:"\s t. (\a. a\s \ a\t \ False) \ s \ t = {}" by auto show "interior s \ interior t = {}" unfolding e_f c_d interior_closed_interval proof(rule *) fix x assume "x\{c<..{e<..i < d\i" "e\i < f\i" "c\i < f\i" "e\i < d\i" unfolding mem_interval using i' + apply-apply(erule_tac[!] x=i in ballE)+ by auto show False using c_d(2)[OF i'] apply- apply(erule_tac disjE) - proof(erule_tac[!] conjE) assume as:"c $$ i = a $$ i" "d $$ i = (a $$ i + b $$ i) / 2" - show False using e_f(2)[of i] and i x unfolding as by(fastforce simp add:field_simps) - next assume as:"c $$ i = (a $$ i + b $$ i) / 2" "d $$ i = b $$ i" - show False using e_f(2)[of i] and i x unfolding as by(fastforce simp add:field_simps) + proof(erule_tac[!] conjE) assume as:"c \ i = a \ i" "d \ i = (a \ i + b \ i) / 2" + show False using e_f(2)[OF i'] and i x unfolding as by(fastforce simp add:field_simps) + next assume as:"c \ i = (a \ i + b \ i) / 2" "d \ i = b \ i" + show False using e_f(2)[OF i'] and i x unfolding as by(fastforce simp add:field_simps) qed qed qed also have "\ ?A = {a..b}" proof(rule set_eqI,rule) fix x assume "x\\?A" then guess Y unfolding Union_iff .. from this(1) guess c unfolding mem_Collect_eq .. then guess d .. note c_d = this[THEN conjunct2,rule_format] `x\Y`[unfolded this[THEN conjunct1]] show "x\{a..b}" unfolding mem_interval proof safe - fix i assume "i x $$ i" "x $$ i \ b $$ i" - using c_d(1)[of i] c_d(2)[unfolded mem_interval,THEN spec[where x=i]] by auto qed + fix i :: 'a assume i: "i\Basis" thus "a \ i \ x \ i" "x \ i \ b \ i" + using c_d(1)[OF i] c_d(2)[unfolded mem_interval,THEN bspec, OF i] by auto qed next fix x assume x:"x\{a..b}" - have "\ic d. (c = a$$i \ d = (a$$i + b$$i) / 2 \ c = (a$$i + b$$i) / 2 \ d = b$$i) \ c\x$$i \ x$$i \ d" - (is "\ic d. ?P i c d") unfolding mem_interval proof(rule,rule) fix i - have "?P i (a$$i) ((a $$ i + b $$ i) / 2) \ ?P i ((a $$ i + b $$ i) / 2) (b$$i)" - using x[unfolded mem_interval,THEN spec[where x=i]] by auto thus "\c d. ?P i c d" by blast - qed thus "x\\?A" unfolding Union_iff unfolding lambda_skolem' unfolding Bex_def mem_Collect_eq + have "\i\Basis. \c d. (c = a\i \ d = (a\i + b\i) / 2 \ c = (a\i + b\i) / 2 \ d = b\i) \ c\x\i \ x\i \ d" + (is "\i\Basis. \c d. ?P i c d") unfolding mem_interval + proof + fix i :: 'a assume i: "i \ Basis" + have "?P i (a\i) ((a \ i + b \ i) / 2) \ ?P i ((a \ i + b \ i) / 2) (b\i)" + using x[unfolded mem_interval,THEN bspec, OF i] by auto thus "\c d. ?P i c d" by blast + qed + thus "x\\?A" + unfolding Union_iff Bex_def mem_Collect_eq choice_Basis_iff apply-apply(erule exE)+ apply(rule_tac x="{xa..xaa}" in exI) unfolding mem_interval by auto qed finally show False using assms by auto qed @@ -1490,8 +1384,8 @@ obtains x where "x \ {a..b}" "\e>0. \c d. x \ {c..d} \ {c..d} \ ball x e \ {c..d} \ {a..b} \ ~P({c..d})" proof- have "\x. \y. \ P {fst x..snd x} \ (\ P {fst y..snd y} \ - (\i fst y$$i \ fst y$$i \ snd y$$i \ snd y$$i \ snd x$$i \ - 2 * (snd y$$i - fst y$$i) \ snd x$$i - fst x$$i))" proof case goal1 thus ?case proof- + (\i\Basis. fst x\i \ fst y\i \ fst y\i \ snd y\i \ snd y\i \ snd x\i \ + 2 * (snd y\i - fst y\i) \ snd x\i - fst x\i))" proof case goal1 thus ?case proof- presume "\ P {fst x..snd x} \ ?thesis" thus ?thesis apply(cases "P {fst x..snd x}") by auto next assume as:"\ P {fst x..snd x}" from interval_bisection_step[of P, OF assms(1-2) as] guess c d . @@ -1499,8 +1393,8 @@ qed qed then guess f apply-apply(drule choice) by(erule exE) note f=this def AB \ "\n. (f ^^ n) (a,b)" def A \ "\n. fst(AB n)" and B \ "\n. snd(AB n)" note ab_def = this AB_def have "A 0 = a" "B 0 = b" "\n. \ P {A(Suc n)..B(Suc n)} \ - (\i A(Suc n)$$i \ A(Suc n)$$i \ B(Suc n)$$i \ B(Suc n)$$i \ B(n)$$i \ - 2 * (B(Suc n)$$i - A(Suc n)$$i) \ B(n)$$i - A(n)$$i)" (is "\n. ?P n") + (\i\Basis. A(n)\i \ A(Suc n)\i \ A(Suc n)\i \ B(Suc n)\i \ B(Suc n)\i \ B(n)\i \ + 2 * (B(Suc n)\i - A(Suc n)\i) \ B(n)\i - A(n)\i)" (is "\n. ?P n") proof- show "A 0 = a" "B 0 = b" unfolding ab_def by auto case goal3 note S = ab_def funpow.simps o_def id_apply show ?case proof(induct n) case 0 thus ?case unfolding S apply(rule f[rule_format]) using assms(3) by auto @@ -1508,29 +1402,28 @@ qed qed note AB = this(1-2) conjunctD2[OF this(3),rule_format] have interv:"\e. 0 < e \ \n. \x\{A n..B n}. \y\{A n..B n}. dist x y < e" - proof- case goal1 guess n using real_arch_pow2[of "(setsum (\i. b$$i - a$$i) {..i. b\i - a\i) Basis) / e"] .. note n=this show ?case apply(rule_tac x=n in exI) proof(rule,rule) fix x y assume xy:"x\{A n..B n}" "y\{A n..B n}" - have "dist x y \ setsum (\i. abs((x - y)$$i)) {.. \ setsum (\i. B n$$i - A n$$i) {..(x - y) $$ i\ \ B n $$ i - A n $$ i" - using xy[unfolded mem_interval,THEN spec[where x=i]] by auto qed - also have "\ \ setsum (\i. b$$i - a$$i) {.. setsum (\i. abs((x - y)\i)) Basis" unfolding dist_norm by(rule norm_le_l1) + also have "\ \ setsum (\i. B n\i - A n\i) Basis" + proof(rule setsum_mono) + fix i :: 'a assume i: "i \ Basis" show "\(x - y) \ i\ \ B n \ i - A n \ i" + using xy[unfolded mem_interval,THEN bspec, OF i] by (auto simp: inner_diff_left) qed + also have "\ \ setsum (\i. b\i - a\i) Basis / 2^n" unfolding setsum_divide_distrib proof(rule setsum_mono) case goal1 thus ?case proof(induct n) case 0 thus ?case unfolding AB by auto - next case (Suc n) have "B (Suc n) $$ i - A (Suc n) $$ i \ (B n $$ i - A n $$ i) / 2" + next case (Suc n) have "B (Suc n) \ i - A (Suc n) \ i \ (B n \ i - A n \ i) / 2" using AB(4)[of i n] using goal1 by auto - also have "\ \ (b $$ i - a $$ i) / 2 ^ Suc n" using Suc by(auto simp add:field_simps) finally show ?case . + also have "\ \ (b \ i - a \ i) / 2 ^ Suc n" using Suc by(auto simp add:field_simps) finally show ?case . qed qed also have "\ < e" using n using goal1 by(auto simp add:field_simps) finally show "dist x y < e" . qed qed - { fix n m ::nat assume "m \ n" then guess d unfolding le_Suc_ex_iff .. note d=this - have "{A n..B n} \ {A m..B m}" unfolding d - proof(induct d) case 0 thus ?case by auto - next case (Suc d) show ?case apply(rule subset_trans[OF _ Suc]) - apply(rule) unfolding mem_interval apply(rule,erule_tac x=i in allE) - proof- case goal1 thus ?case using AB(4)[of i "m + d"] by(auto simp add:field_simps) - qed qed } note ABsubset = this + { fix n m :: nat assume "m \ n" then have "{A n..B n} \ {A m..B m}" + proof(induct rule: inc_induct) + case (step i) show ?case + using AB(4) by (intro order_trans[OF step(2)] subset_interval_imp) auto + qed simp } note ABsubset = this have "\a. \n. a\{A n..B n}" apply(rule decreasing_closed_nest[rule_format,OF closed_interval _ ABsubset interv]) proof- fix n show "{A n..B n} \ {}" apply(cases "0 'm::ordered_euclidean_space" - assumes "f integrable_on s" shows "integral s (\x. f x $$ k) = integral s f $$ k" + assumes "f integrable_on s" shows "integral s (\x. f x \ k) = integral s f \ k" unfolding integral_linear[OF assms(1) bounded_linear_component,unfolded o_def] .. lemma has_integral_setsum: @@ -1852,8 +1745,9 @@ apply(rule integral_unique) using has_integral_empty . lemma has_integral_refl[intro]: shows "(f has_integral 0) {a..a}" "(f has_integral 0) {a::'a::ordered_euclidean_space}" -proof- have *:"{a} = {a..a}" apply(rule set_eqI) unfolding mem_interval singleton_iff euclidean_eq[where 'a='a] - apply safe prefer 3 apply(erule_tac x=i in allE) by(auto simp add: field_simps) +proof- + have *:"{a} = {a..a}" apply(rule set_eqI) unfolding mem_interval singleton_iff euclidean_eq_iff[where 'a='a] + apply safe prefer 3 apply(erule_tac x=b in ballE) by(auto simp add: field_simps) show "(f has_integral 0) {a..a}" "(f has_integral 0) {a}" unfolding * apply(rule_tac[!] has_integral_null) unfolding content_eq_0_interior unfolding interior_closed_interval using interval_sing by auto qed @@ -1913,42 +1807,48 @@ subsection {* Additivity of integral on abutting intervals. *} -lemma interval_split: fixes a::"'a::ordered_euclidean_space" assumes "k {x. x$$k \ c} = {a .. (\\ i. if i = k then min (b$$k) c else b$$i)}" - "{a..b} \ {x. x$$k \ c} = {(\\ i. if i = k then max (a$$k) c else a$$i) .. b}" - apply(rule_tac[!] set_eqI) unfolding Int_iff mem_interval mem_Collect_eq using assms by auto - -lemma content_split: fixes a::"'a::ordered_euclidean_space" assumes "k {x. x$$k \ c}) + content({a..b} \ {x. x$$k >= c})" -proof- note simps = interval_split[OF assms] content_closed_interval_cases eucl_le[where 'a='a] - { presume "a\b \ ?thesis" thus ?thesis apply(cases "a\b") unfolding simps using assms by auto } - have *:"{..x. finite ({..x. x\{.. Basis" + shows + "{a..b} \ {x. x\k \ c} = {a .. (\i\Basis. (if i = k then min (b\k) c else b\i) *\<^sub>R i)}" + "{a..b} \ {x. x\k \ c} = {(\i\Basis. (if i = k then max (a\k) c else a\i) *\<^sub>R i) .. b}" + apply(rule_tac[!] set_eqI) unfolding Int_iff mem_interval mem_Collect_eq using assms + by auto + +lemma content_split: fixes a::"'a::ordered_euclidean_space" assumes "k\Basis" shows + "content {a..b} = content({a..b} \ {x. x\k \ c}) + content({a..b} \ {x. x\k >= c})" +proof cases + note simps = interval_split[OF assms] content_closed_interval_cases eucl_le[where 'a='a] + have *:"Basis = insert k (Basis - {k})" "\x. finite (Basis-{x})" "\x. x\Basis-{x}" using assms by auto - have *:"\X Y Z. (\i\{..i\{..i\{..i\{..X Y Z. (\i\Basis. Z i (if i = k then X else Y i)) = Z k X * (\i\Basis-{k}. Z i (Y i))" + "(\i\Basis. b\i - a\i) = (\i\Basis-{k}. b\i - a\i) * (b\k - a\k)" apply(subst *(1)) defer apply(subst *(1)) unfolding setprod_insert[OF *(2-)] by auto - assume as:"a\b" moreover have "\x. min (b $$ k) c = max (a $$ k) c - \ x* (b$$k - a$$k) = x*(max (a $$ k) c - a $$ k) + x*(b $$ k - max (a $$ k) c)" + assume as:"a\b" moreover have "\x. min (b \ k) c = max (a \ k) c + \ x* (b\k - a\k) = x*(max (a \ k) c - a \ k) + x*(b \ k - max (a \ k) c)" by (auto simp add:field_simps) - moreover have **:"(\i\ i. if i = k then min (b $$ k) c else b $$ i)::'a) $$ i - a $$ i) = - (\ii\ i. if i = k then max (a $$ k) c else a $$ i)::'a) $$ i) = - (\i a $$ k \ c \ \ c \ b $$ k \ False" + moreover have **:"(\i\Basis. ((\i\Basis. (if i = k then min (b \ k) c else b \ i) *\<^sub>R i) \ i - a \ i)) = + (\i\Basis. (if i = k then min (b \ k) c else b \ i) - a \ i)" + "(\i\Basis. b \ i - ((\i\Basis. (if i = k then max (a \ k) c else a \ i) *\<^sub>R i) \ i)) = + (\i\Basis. b \ i - (if i = k then max (a \ k) c else a \ i))" + by (auto intro!: setprod_cong) + have "\ a \ k \ c \ \ c \ b \ k \ False" unfolding not_le using as[unfolded eucl_le[where 'a='a],rule_format,of k] assms by auto ultimately show ?thesis using assms unfolding simps ** - unfolding *(1)[of "\i x. b$$i - x"] *(1)[of "\i x. x - a$$i"] unfolding *(2) - apply(subst(2) euclidean_lambda_beta''[where 'a='a]) - apply(subst(3) euclidean_lambda_beta''[where 'a='a]) by auto + unfolding *(1)[of "\i x. b\i - x"] *(1)[of "\i x. x - a\i"] unfolding *(2) + by auto +next + assume "\ a \ b" then have "{a .. b} = {}" + unfolding interval_eq_empty by (auto simp: eucl_le[where 'a='a] not_le) + then show ?thesis by auto qed lemma division_split_left_inj: fixes type::"'a::ordered_euclidean_space" assumes "d division_of i" "k1 \ d" "k2 \ d" "k1 \ k2" - "k1 \ {x::'a. x$$k \ c} = k2 \ {x. x$$k \ c}"and k:"k {x. x$$k \ c}) = 0" + "k1 \ {x::'a. x\k \ c} = k2 \ {x. x\k \ c}"and k:"k\Basis" + shows "content(k1 \ {x. x\k \ c}) = 0" proof- note d=division_ofD[OF assms(1)] - have *:"\a b::'a. \ c. (content({a..b} \ {x. x$$k \ c}) = 0 \ interior({a..b} \ {x. x$$k \ c}) = {})" + have *:"\a b::'a. \ c. (content({a..b} \ {x. x\k \ c}) = 0 \ interior({a..b} \ {x. x\k \ c}) = {})" unfolding interval_split[OF k] content_eq_0_interior by auto guess u1 v1 using d(4)[OF assms(2)] apply-by(erule exE)+ note uv1=this guess u2 v2 using d(4)[OF assms(3)] apply-by(erule exE)+ note uv2=this @@ -1958,10 +1858,10 @@ lemma division_split_right_inj: fixes type::"'a::ordered_euclidean_space" assumes "d division_of i" "k1 \ d" "k2 \ d" "k1 \ k2" - "k1 \ {x::'a. x$$k \ c} = k2 \ {x. x$$k \ c}" and k:"k {x. x$$k \ c}) = 0" + "k1 \ {x::'a. x\k \ c} = k2 \ {x. x\k \ c}" and k:"k\Basis" + shows "content(k1 \ {x. x\k \ c}) = 0" proof- note d=division_ofD[OF assms(1)] - have *:"\a b::'a. \ c. (content({a..b} \ {x. x$$k >= c}) = 0 \ interior({a..b} \ {x. x$$k >= c}) = {})" + have *:"\a b::'a. \ c. (content({a..b} \ {x. x\k >= c}) = 0 \ interior({a..b} \ {x. x\k >= c}) = {})" unfolding interval_split[OF k] content_eq_0_interior by auto guess u1 v1 using d(4)[OF assms(2)] apply-by(erule exE)+ note uv1=this guess u2 v2 using d(4)[OF assms(3)] apply-by(erule exE)+ note uv2=this @@ -1970,25 +1870,25 @@ defer apply(subst assms(5)[unfolded uv1 uv2]) unfolding uv1 uv2 by auto qed lemma tagged_division_split_left_inj: fixes x1::"'a::ordered_euclidean_space" - assumes "d tagged_division_of i" "(x1,k1) \ d" "(x2,k2) \ d" "k1 \ k2" "k1 \ {x. x$$k \ c} = k2 \ {x. x$$k \ c}" - and k:"k {x. x$$k \ c}) = 0" + assumes "d tagged_division_of i" "(x1,k1) \ d" "(x2,k2) \ d" "k1 \ k2" "k1 \ {x. x\k \ c} = k2 \ {x. x\k \ c}" + and k:"k\Basis" + shows "content(k1 \ {x. x\k \ c}) = 0" proof- have *:"\a b c. (a,b) \ c \ b \ snd ` c" unfolding image_iff apply(rule_tac x="(a,b)" in bexI) by auto show ?thesis apply(rule division_split_left_inj[OF division_of_tagged_division[OF assms(1)]]) apply(rule_tac[1-2] *) using assms(2-) by auto qed lemma tagged_division_split_right_inj: fixes x1::"'a::ordered_euclidean_space" - assumes "d tagged_division_of i" "(x1,k1) \ d" "(x2,k2) \ d" "k1 \ k2" "k1 \ {x. x$$k \ c} = k2 \ {x. x$$k \ c}" - and k:"k {x. x$$k \ c}) = 0" + assumes "d tagged_division_of i" "(x1,k1) \ d" "(x2,k2) \ d" "k1 \ k2" "k1 \ {x. x\k \ c} = k2 \ {x. x\k \ c}" + and k:"k\Basis" + shows "content(k1 \ {x. x\k \ c}) = 0" proof- have *:"\a b c. (a,b) \ c \ b \ snd ` c" unfolding image_iff apply(rule_tac x="(a,b)" in bexI) by auto show ?thesis apply(rule division_split_right_inj[OF division_of_tagged_division[OF assms(1)]]) apply(rule_tac[1-2] *) using assms(2-) by auto qed lemma division_split: fixes a::"'a::ordered_euclidean_space" - assumes "p division_of {a..b}" and k:"k {x. x$$k \ c} | l. l \ p \ ~(l \ {x. x$$k \ c} = {})} division_of({a..b} \ {x. x$$k \ c})" (is "?p1 division_of ?I1") and - "{l \ {x. x$$k \ c} | l. l \ p \ ~(l \ {x. x$$k \ c} = {})} division_of ({a..b} \ {x. x$$k \ c})" (is "?p2 division_of ?I2") + assumes "p division_of {a..b}" and k:"k\Basis" + shows "{l \ {x. x\k \ c} | l. l \ p \ ~(l \ {x. x\k \ c} = {})} division_of({a..b} \ {x. x\k \ c})" (is "?p1 division_of ?I1") and + "{l \ {x. x\k \ c} | l. l \ p \ ~(l \ {x. x\k \ c} = {})} division_of ({a..b} \ {x. x\k \ c})" (is "?p2 division_of ?I2") proof(rule_tac[!] division_ofI) note p=division_ofD[OF assms(1)] show "finite ?p1" "finite ?p2" using p(1) by auto show "\?p1 = ?I1" "\?p2 = ?I2" unfolding p(6)[THEN sym] by auto { fix k assume "k\?p1" then guess l unfolding mem_Collect_eq apply-by(erule exE,(erule conjE)+) note l=this @@ -2006,34 +1906,34 @@ qed lemma has_integral_split: fixes f::"'a::ordered_euclidean_space \ 'b::real_normed_vector" - assumes "(f has_integral i) ({a..b} \ {x. x$$k \ c})" "(f has_integral j) ({a..b} \ {x. x$$k \ c})" and k:"k {x. x\k \ c})" "(f has_integral j) ({a..b} \ {x. x\k \ c})" and k:"k\Basis" shows "(f has_integral (i + j)) ({a..b})" proof(unfold has_integral,rule,rule) case goal1 hence e:"e/2>0" by auto guess d1 using has_integralD[OF assms(1)[unfolded interval_split[OF k]] e] . note d1=this[unfolded interval_split[THEN sym,OF k]] guess d2 using has_integralD[OF assms(2)[unfolded interval_split[OF k]] e] . note d2=this[unfolded interval_split[THEN sym,OF k]] - let ?d = "\x. if x$$k = c then (d1 x \ d2 x) else ball x (abs(x$$k - c)) \ d1 x \ d2 x" + let ?d = "\x. if x\k = c then (d1 x \ d2 x) else ball x (abs(x\k - c)) \ d1 x \ d2 x" show ?case apply(rule_tac x="?d" in exI,rule) defer apply(rule,rule,(erule conjE)+) proof- show "gauge ?d" using d1(1) d2(1) unfolding gauge_def by auto fix p assume "p tagged_division_of {a..b}" "?d fine p" note p = this tagged_division_ofD[OF this(1)] - have lem0:"\x kk. (x,kk) \ p \ ~(kk \ {x. x$$k \ c} = {}) \ x$$k \ c" - "\x kk. (x,kk) \ p \ ~(kk \ {x. x$$k \ c} = {}) \ x$$k \ c" + have lem0:"\x kk. (x,kk) \ p \ ~(kk \ {x. x\k \ c} = {}) \ x\k \ c" + "\x kk. (x,kk) \ p \ ~(kk \ {x. x\k \ c} = {}) \ x\k \ c" proof- fix x kk assume as:"(x,kk)\p" - show "~(kk \ {x. x$$k \ c} = {}) \ x$$k \ c" + show "~(kk \ {x. x\k \ c} = {}) \ x\k \ c" proof(rule ccontr) case goal1 - from this(2)[unfolded not_le] have "kk \ ball x \x $$ k - c\" + from this(2)[unfolded not_le] have "kk \ ball x \x \ k - c\" using p(2)[unfolded fine_def,rule_format,OF as,unfolded split_conv] by auto - hence "\y. y \ ball x \x $$ k - c\ \ {x. x $$ k \ c}" using goal1(1) by blast - then guess y .. hence "\x $$ k - y $$ k\ < \x $$ k - c\" "y$$k \ c" apply-apply(rule le_less_trans) - using component_le_norm[of "x - y" k] by(auto simp add:dist_norm) + hence "\y. y \ ball x \x \ k - c\ \ {x. x \ k \ c}" using goal1(1) by blast + then guess y .. hence "\x \ k - y \ k\ < \x \ k - c\" "y\k \ c" apply-apply(rule le_less_trans) + using Basis_le_norm[OF k, of "x - y"] by (auto simp add: dist_norm inner_diff_left) thus False using goal1(2)[unfolded not_le] by(auto simp add:field_simps) qed - show "~(kk \ {x. x$$k \ c} = {}) \ x$$k \ c" + show "~(kk \ {x. x\k \ c} = {}) \ x\k \ c" proof(rule ccontr) case goal1 - from this(2)[unfolded not_le] have "kk \ ball x \x $$ k - c\" + from this(2)[unfolded not_le] have "kk \ ball x \x \ k - c\" using p(2)[unfolded fine_def,rule_format,OF as,unfolded split_conv] by auto - hence "\y. y \ ball x \x $$ k - c\ \ {x. x $$ k \ c}" using goal1(1) by blast - then guess y .. hence "\x $$ k - y $$ k\ < \x $$ k - c\" "y$$k \ c" apply-apply(rule le_less_trans) - using component_le_norm[of "x - y" k] by(auto simp add:dist_norm) + hence "\y. y \ ball x \x \ k - c\ \ {x. x \ k \ c}" using goal1(1) by blast + then guess y .. hence "\x \ k - y \ k\ < \x \ k - c\" "y\k \ c" apply-apply(rule le_less_trans) + using Basis_le_norm[OF k, of "x - y"] by (auto simp add: dist_norm inner_diff_left) thus False using goal1(2)[unfolded not_le] by(auto simp add:field_simps) qed qed @@ -2053,15 +1953,15 @@ qed auto have lem4:"\g. (\(x,l). content (g l) *\<^sub>R f x) = (\(x,l). content l *\<^sub>R f x) o (\(x,l). (x,g l))" apply(rule ext) by auto - let ?M1 = "{(x,kk \ {x. x$$k \ c}) |x kk. (x,kk) \ p \ kk \ {x. x$$k \ c} \ {}}" + let ?M1 = "{(x,kk \ {x. x\k \ c}) |x kk. (x,kk) \ p \ kk \ {x. x\k \ c} \ {}}" have "norm ((\(x, k)\?M1. content k *\<^sub>R f x) - i) < e/2" apply(rule d1(2),rule tagged_division_ofI) apply(rule lem2 p(3))+ prefer 6 apply(rule fineI) - proof- show "\{k. \x. (x, k) \ ?M1} = {a..b} \ {x. x$$k \ c}" unfolding p(8)[THEN sym] by auto + proof- show "\{k. \x. (x, k) \ ?M1} = {a..b} \ {x. x\k \ c}" unfolding p(8)[THEN sym] by auto fix x l assume xl:"(x,l)\?M1" then guess x' l' unfolding mem_Collect_eq apply- unfolding Pair_eq apply((erule exE)+,(erule conjE)+) . note xl'=this have "l' \ d1 x'" apply(rule order_trans[OF fineD[OF p(2) xl'(3)]]) by auto thus "l \ d1 x" unfolding xl' by auto - show "x\l" "l \ {a..b} \ {x. x $$ k \ c}" unfolding xl' using p(4-6)[OF xl'(3)] using xl'(4) + show "x\l" "l \ {a..b} \ {x. x \ k \ c}" unfolding xl' using p(4-6)[OF xl'(3)] using xl'(4) using lem0(1)[OF xl'(3-4)] by auto show "\a b. l = {a..b}" unfolding xl' using p(6)[OF xl'(3)] by(fastforce simp add: interval_split[OF k,where c=c]) fix y r let ?goal = "interior l \ interior r = {}" assume yr:"(y,r)\?M1" @@ -2073,15 +1973,15 @@ thus ?thesis using p(7)[OF xl'(3) yr'(3)] using as unfolding xl' yr' by auto qed qed moreover - let ?M2 = "{(x,kk \ {x. x$$k \ c}) |x kk. (x,kk) \ p \ kk \ {x. x$$k \ c} \ {}}" + let ?M2 = "{(x,kk \ {x. x\k \ c}) |x kk. (x,kk) \ p \ kk \ {x. x\k \ c} \ {}}" have "norm ((\(x, k)\?M2. content k *\<^sub>R f x) - j) < e/2" apply(rule d2(2),rule tagged_division_ofI) apply(rule lem2 p(3))+ prefer 6 apply(rule fineI) - proof- show "\{k. \x. (x, k) \ ?M2} = {a..b} \ {x. x$$k \ c}" unfolding p(8)[THEN sym] by auto + proof- show "\{k. \x. (x, k) \ ?M2} = {a..b} \ {x. x\k \ c}" unfolding p(8)[THEN sym] by auto fix x l assume xl:"(x,l)\?M2" then guess x' l' unfolding mem_Collect_eq apply- unfolding Pair_eq apply((erule exE)+,(erule conjE)+) . note xl'=this have "l' \ d2 x'" apply(rule order_trans[OF fineD[OF p(2) xl'(3)]]) by auto thus "l \ d2 x" unfolding xl' by auto - show "x\l" "l \ {a..b} \ {x. x $$ k \ c}" unfolding xl' using p(4-6)[OF xl'(3)] using xl'(4) + show "x\l" "l \ {a..b} \ {x. x \ k \ c}" unfolding xl' using p(4-6)[OF xl'(3)] using xl'(4) using lem0(2)[OF xl'(3-4)] by auto show "\a b. l = {a..b}" unfolding xl' using p(6)[OF xl'(3)] by(fastforce simp add: interval_split[OF k, where c=c]) fix y r let ?goal = "interior l \ interior r = {}" assume yr:"(y,r)\?M2" @@ -2098,98 +1998,95 @@ also { have *:"\x y. x = (0::real) \ x *\<^sub>R (y::'b) = 0" using scaleR_zero_left by auto have "((\(x, k)\?M1. content k *\<^sub>R f x) - i) + ((\(x, k)\?M2. content k *\<^sub>R f x) - j) = (\(x, k)\?M1. content k *\<^sub>R f x) + (\(x, k)\?M2. content k *\<^sub>R f x) - (i + j)" by auto - also have "\ = (\(x, ka)\p. content (ka \ {x. x $$ k \ c}) *\<^sub>R f x) + - (\(x, ka)\p. content (ka \ {x. c \ x $$ k}) *\<^sub>R f x) - (i + j)" + also have "\ = (\(x, ka)\p. content (ka \ {x. x \ k \ c}) *\<^sub>R f x) + + (\(x, ka)\p. content (ka \ {x. c \ x \ k}) *\<^sub>R f x) - (i + j)" unfolding lem3[OF p(3)] apply(subst setsum_reindex_nonzero[OF p(3)]) defer apply(subst setsum_reindex_nonzero[OF p(3)]) defer unfolding lem4[THEN sym] apply(rule refl) unfolding split_paired_all split_conv apply(rule_tac[!] *) proof- case goal1 thus ?case apply- apply(rule tagged_division_split_left_inj [OF p(1), of a b aa ba]) using k by auto next case goal2 thus ?case apply- apply(rule tagged_division_split_right_inj[OF p(1), of a b aa ba]) using k by auto qed also note setsum_addf[THEN sym] - also have *:"\x. x\p \ (\(x, ka). content (ka \ {x. x $$ k \ c}) *\<^sub>R f x) x + (\(x, ka). content (ka \ {x. c \ x $$ k}) *\<^sub>R f x) x + also have *:"\x. x\p \ (\(x, ka). content (ka \ {x. x \ k \ c}) *\<^sub>R f x) x + (\(x, ka). content (ka \ {x. c \ x \ k}) *\<^sub>R f x) x = (\(x,ka). content ka *\<^sub>R f x) x" unfolding split_paired_all split_conv proof- fix a b assume "(a,b) \ p" from p(6)[OF this] guess u v apply-by(erule exE)+ note uv=this - thus "content (b \ {x. x $$ k \ c}) *\<^sub>R f a + content (b \ {x. c \ x $$ k}) *\<^sub>R f a = content b *\<^sub>R f a" + thus "content (b \ {x. x \ k \ c}) *\<^sub>R f a + content (b \ {x. c \ x \ k}) *\<^sub>R f a = content b *\<^sub>R f a" unfolding scaleR_left_distrib[THEN sym] unfolding uv content_split[OF k,of u v c] by auto qed note setsum_cong2[OF this] - finally have "(\(x, k)\{(x, kk \ {x. x $$ k \ c}) |x kk. (x, kk) \ p \ kk \ {x. x $$ k \ c} \ {}}. content k *\<^sub>R f x) - i + - ((\(x, k)\{(x, kk \ {x. c \ x $$ k}) |x kk. (x, kk) \ p \ kk \ {x. c \ x $$ k} \ {}}. content k *\<^sub>R f x) - j) = + finally have "(\(x, k)\{(x, kk \ {x. x \ k \ c}) |x kk. (x, kk) \ p \ kk \ {x. x \ k \ c} \ {}}. content k *\<^sub>R f x) - i + + ((\(x, k)\{(x, kk \ {x. c \ x \ k}) |x kk. (x, kk) \ p \ kk \ {x. c \ x \ k} \ {}}. content k *\<^sub>R f x) - j) = (\(x, ka)\p. content ka *\<^sub>R f x) - (i + j)" by auto } finally show "norm ((\(x, k)\p. content k *\<^sub>R f x) - (i + j)) < e" by auto qed qed -(*lemma has_integral_split_cart: fixes f::"real^'n \ 'a::real_normed_vector" - assumes "(f has_integral i) ({a..b} \ {x. x$k \ c})" "(f has_integral j) ({a..b} \ {x. x$k \ c})" - shows "(f has_integral (i + j)) ({a..b})" *) - subsection {* A sort of converse, integrability on subintervals. *} lemma tagged_division_union_interval: fixes a::"'a::ordered_euclidean_space" - assumes "p1 tagged_division_of ({a..b} \ {x. x$$k \ (c::real)})" "p2 tagged_division_of ({a..b} \ {x. x$$k \ c})" - and k:"k {x. x\k \ (c::real)})" "p2 tagged_division_of ({a..b} \ {x. x\k \ c})" + and k:"k\Basis" shows "(p1 \ p2) tagged_division_of ({a..b})" -proof- have *:"{a..b} = ({a..b} \ {x. x$$k \ c}) \ ({a..b} \ {x. x$$k \ c})" by auto +proof- have *:"{a..b} = ({a..b} \ {x. x\k \ c}) \ ({a..b} \ {x. x\k \ c})" by auto show ?thesis apply(subst *) apply(rule tagged_division_union[OF assms(1-2)]) unfolding interval_split[OF k] interior_closed_interval using k - by(auto simp add: eucl_less[where 'a='a] elim!:allE[where x=k]) qed + by(auto simp add: eucl_less[where 'a='a] elim!: ballE[where x=k]) qed lemma has_integral_separate_sides: fixes f::"'a::ordered_euclidean_space \ 'b::real_normed_vector" - assumes "(f has_integral i) ({a..b})" "e>0" and k:"kp1 p2. p1 tagged_division_of ({a..b} \ {x. x$$k \ c}) \ d fine p1 \ - p2 tagged_division_of ({a..b} \ {x. x$$k \ c}) \ d fine p2 + assumes "(f has_integral i) ({a..b})" "e>0" and k:"k\Basis" + obtains d where "gauge d" "(\p1 p2. p1 tagged_division_of ({a..b} \ {x. x\k \ c}) \ d fine p1 \ + p2 tagged_division_of ({a..b} \ {x. x\k \ c}) \ d fine p2 \ norm((setsum (\(x,k). content k *\<^sub>R f x) p1 + setsum (\(x,k). content k *\<^sub>R f x) p2) - i) < e)" proof- guess d using has_integralD[OF assms(1-2)] . note d=this show ?thesis apply(rule that[of d]) apply(rule d) apply(rule,rule,rule,(erule conjE)+) - proof- fix p1 p2 assume "p1 tagged_division_of {a..b} \ {x. x $$ k \ c}" "d fine p1" note p1=tagged_division_ofD[OF this(1)] this - assume "p2 tagged_division_of {a..b} \ {x. c \ x $$ k}" "d fine p2" note p2=tagged_division_ofD[OF this(1)] this + proof- fix p1 p2 assume "p1 tagged_division_of {a..b} \ {x. x \ k \ c}" "d fine p1" note p1=tagged_division_ofD[OF this(1)] this + assume "p2 tagged_division_of {a..b} \ {x. c \ x \ k}" "d fine p2" note p2=tagged_division_ofD[OF this(1)] this note tagged_division_union_interval[OF p1(7) p2(7)] note p12 = tagged_division_ofD[OF this] this have "norm ((\(x, k)\p1. content k *\<^sub>R f x) + (\(x, k)\p2. content k *\<^sub>R f x) - i) = norm ((\(x, k)\p1 \ p2. content k *\<^sub>R f x) - i)" apply(subst setsum_Un_zero) apply(rule p1 p2)+ apply(rule) unfolding split_paired_all split_conv proof- fix a b assume ab:"(a,b) \ p1 \ p2" have "(a,b) \ p1" using ab by auto from p1(4)[OF this] guess u v apply-by(erule exE)+ note uv =this - have "b \ {x. x$$k = c}" using ab p1(3)[of a b] p2(3)[of a b] by fastforce - moreover have "interior {x::'a. x $$ k = c} = {}" - proof(rule ccontr) case goal1 then obtain x where x:"x\interior {x::'a. x$$k = c}" by auto + have "b \ {x. x\k = c}" using ab p1(3)[of a b] p2(3)[of a b] by fastforce + moreover have "interior {x::'a. x \ k = c} = {}" + proof(rule ccontr) case goal1 then obtain x where x:"x\interior {x::'a. x\k = c}" by auto then guess e unfolding mem_interior .. note e=this - have x:"x$$k = c" using x interior_subset by fastforce - have *:"\i. i \(x - (x + (\\ i. if i = k then e / 2 else 0))) $$ i\ - = (if i = k then e/2 else 0)" using e by auto - have "(\i(x - (x + (\\ i. if i = k then e / 2 else 0))) $$ i\) = - (\ik = c" using x interior_subset by fastforce + have *:"\i. i\Basis \ \(x - (x + (e / 2) *\<^sub>R k)) \ i\ + = (if i = k then e/2 else 0)" using e k by (auto simp: inner_simps inner_not_same_Basis) + have "(\i\Basis. \(x - (x + (e / 2 ) *\<^sub>R k)) \ i\) = + (\i\Basis. (if i = k then e / 2 else 0))" apply(rule setsum_cong2) apply(subst *) by auto also have "... < e" apply(subst setsum_delta) using e by auto - finally have "x + (\\ i. if i = k then e/2 else 0) \ ball x e" unfolding mem_ball dist_norm - by(rule le_less_trans[OF norm_le_l1]) - hence "x + (\\ i. if i = k then e/2 else 0) \ {x. x$$k = c}" using e by auto - thus False unfolding mem_Collect_eq using e x k by auto + finally have "x + (e/2) *\<^sub>R k \ ball x e" + unfolding mem_ball dist_norm by(rule le_less_trans[OF norm_le_l1]) + hence "x + (e/2) *\<^sub>R k \ {x. x\k = c}" using e by auto + thus False unfolding mem_Collect_eq using e x k by (auto simp: inner_simps) qed ultimately have "content b = 0" unfolding uv content_eq_0_interior apply-apply(drule interior_mono) by auto thus "content b *\<^sub>R f a = 0" by auto qed auto also have "\ < e" by(rule k d(2) p12 fine_union p1 p2)+ finally show "norm ((\(x, k)\p1. content k *\<^sub>R f x) + (\(x, k)\p2. content k *\<^sub>R f x) - i) < e" . qed qed -lemma integrable_split[intro]: fixes f::"'a::ordered_euclidean_space \ 'b::{real_normed_vector,complete_space}" - assumes "f integrable_on {a..b}" and k:"k {x. x$$k \ c})" (is ?t1) and "f integrable_on ({a..b} \ {x. x$$k \ c})" (is ?t2) +lemma integrable_split[intro]: + fixes f::"'a::ordered_euclidean_space \ 'b::{real_normed_vector,complete_space}" + assumes "f integrable_on {a..b}" and k:"k\Basis" + shows "f integrable_on ({a..b} \ {x. x\k \ c})" (is ?t1) and "f integrable_on ({a..b} \ {x. x\k \ c})" (is ?t2) proof- guess y using assms(1) unfolding integrable_on_def .. note y=this - def b' \ "(\\ i. if i = k then min (b$$k) c else b$$i)::'a" - and a' \ "(\\ i. if i = k then max (a$$k) c else a$$i)::'a" + def b' \ "\i\Basis. (if i = k then min (b\k) c else b\i)*\<^sub>R i::'a" + def a' \ "\i\Basis. (if i = k then max (a\k) c else a\i)*\<^sub>R i::'a" show ?t1 ?t2 unfolding interval_split[OF k] integrable_cauchy unfolding interval_split[THEN sym,OF k] proof(rule_tac[!] allI impI)+ fix e::real assume "e>0" hence "e/2>0" by auto from has_integral_separate_sides[OF y this k,of c] guess d . note d=this[rule_format] let ?P = "\A. \d. gauge d \ (\p1 p2. p1 tagged_division_of {a..b} \ A \ d fine p1 \ p2 tagged_division_of {a..b} \ A \ d fine p2 \ norm ((\(x, k)\p1. content k *\<^sub>R f x) - (\(x, k)\p2. content k *\<^sub>R f x)) < e)" - show "?P {x. x $$ k \ c}" apply(rule_tac x=d in exI) apply(rule,rule d) apply(rule,rule,rule) - proof- fix p1 p2 assume as:"p1 tagged_division_of {a..b} \ {x. x $$ k \ c} \ d fine p1 - \ p2 tagged_division_of {a..b} \ {x. x $$ k \ c} \ d fine p2" + show "?P {x. x \ k \ c}" apply(rule_tac x=d in exI) apply(rule,rule d) apply(rule,rule,rule) + proof- fix p1 p2 assume as:"p1 tagged_division_of {a..b} \ {x. x \ k \ c} \ d fine p1 + \ p2 tagged_division_of {a..b} \ {x. x \ k \ c} \ d fine p2" show "norm ((\(x, k)\p1. content k *\<^sub>R f x) - (\(x, k)\p2. content k *\<^sub>R f x)) < e" proof- guess p using fine_division_exists[OF d(1), of a' b] . note p=this show ?thesis using norm_triangle_half_l[OF d(2)[of p1 p] d(2)[of p2 p]] using as unfolding interval_split[OF k] b'_def[symmetric] a'_def[symmetric] using p using assms by(auto simp add:algebra_simps) qed qed - show "?P {x. x $$ k \ c}" apply(rule_tac x=d in exI) apply(rule,rule d) apply(rule,rule,rule) - proof- fix p1 p2 assume as:"p1 tagged_division_of {a..b} \ {x. x $$ k \ c} \ d fine p1 - \ p2 tagged_division_of {a..b} \ {x. x $$ k \ c} \ d fine p2" + show "?P {x. x \ k \ c}" apply(rule_tac x=d in exI) apply(rule,rule d) apply(rule,rule,rule) + proof- fix p1 p2 assume as:"p1 tagged_division_of {a..b} \ {x. x \ k \ c} \ d fine p1 + \ p2 tagged_division_of {a..b} \ {x. x \ k \ c} \ d fine p2" show "norm ((\(x, k)\p1. content k *\<^sub>R f x) - (\(x, k)\p2. content k *\<^sub>R f x)) < e" proof- guess p using fine_division_exists[OF d(1), of a b'] . note p=this show ?thesis using norm_triangle_half_l[OF d(2)[of p p1] d(2)[of p p2]] @@ -2203,13 +2100,13 @@ definition operative :: "('a \ 'a \ 'a) \ (('b::ordered_euclidean_space) set \ 'a) \ bool" where "operative opp f \ (\a b. content {a..b} = 0 \ f {a..b} = neutral(opp)) \ - (\a b c. \k {x. x$$k \ c})) - (f({a..b} \ {x. x$$k \ c})))" + (\a b c. \k\Basis. f({a..b}) = + opp (f({a..b} \ {x. x\k \ c})) + (f({a..b} \ {x. x\k \ c})))" lemma operativeD[dest]: fixes type::"'a::ordered_euclidean_space" assumes "operative opp f" shows "\a b. content {a..b} = 0 \ f {a..b::'a} = neutral(opp)" - "\a b c k. k f({a..b}) = opp (f({a..b} \ {x. x$$k \ c})) (f({a..b} \ {x. x$$k \ c}))" + "\a b c k. k\Basis \ f({a..b}) = opp (f({a..b} \ {x. x\k \ c})) (f({a..b} \ {x. x\k \ c}))" using assms unfolding operative_def by auto lemma operative_trivial: @@ -2342,18 +2239,20 @@ lemma operative_integral: fixes f::"'a::ordered_euclidean_space \ 'b::banach" shows "operative (lifted(op +)) (\i. if f integrable_on i then Some(integral i f) else None)" unfolding operative_def unfolding neutral_lifted[OF monoidal_monoid] neutral_add - apply(rule,rule,rule,rule) defer apply(rule allI impI)+ -proof- fix a b c k assume k:"k {x. x $$ k \ c} then Some (integral ({a..b} \ {x. x $$ k \ c}) f) else None) - (if f integrable_on {a..b} \ {x. c \ x $$ k} then Some (integral ({a..b} \ {x. c \ x $$ k}) f) else None)" + apply(rule,rule,rule,rule) defer apply(rule allI ballI)+ +proof- + fix a b c and k :: 'a assume k:"k\Basis" + show "(if f integrable_on {a..b} then Some (integral {a..b} f) else None) = + lifted op + (if f integrable_on {a..b} \ {x. x \ k \ c} then Some (integral ({a..b} \ {x. x \ k \ c}) f) else None) + (if f integrable_on {a..b} \ {x. c \ x \ k} then Some (integral ({a..b} \ {x. c \ x \ k}) f) else None)" proof(cases "f integrable_on {a..b}") case True show ?thesis unfolding if_P[OF True] using k apply- unfolding if_P[OF integrable_split(1)[OF True]] unfolding if_P[OF integrable_split(2)[OF True]] unfolding lifted.simps option.inject apply(rule integral_unique) apply(rule has_integral_split[OF _ _ k]) apply(rule_tac[!] integrable_integral integrable_split)+ using True k by auto - next case False have "(\ (f integrable_on {a..b} \ {x. x $$ k \ c})) \ (\ ( f integrable_on {a..b} \ {x. c \ x $$ k}))" + next case False have "(\ (f integrable_on {a..b} \ {x. x \ k \ c})) \ (\ ( f integrable_on {a..b} \ {x. c \ x \ k}))" proof(rule ccontr) case goal1 hence "f integrable_on {a..b}" apply- unfolding integrable_on_def - apply(rule_tac x="integral ({a..b} \ {x. x $$ k \ c}) f + integral ({a..b} \ {x. x $$ k \ c}) f" in exI) + apply(rule_tac x="integral ({a..b} \ {x. x \ k \ c}) f + integral ({a..b} \ {x. x \ k \ c}) f" in exI) apply(rule has_integral_split[OF _ _ k]) apply(rule_tac[!] integrable_integral) by auto thus False using False by auto qed thus ?thesis using False by auto @@ -2365,91 +2264,110 @@ subsection {* Points of division of a partition. *} definition "division_points (k::('a::ordered_euclidean_space) set) d = - {(j,x). j (interval_lowerbound k)$$j < x \ x < (interval_upperbound k)$$j \ - (\i\d. (interval_lowerbound i)$$j = x \ (interval_upperbound i)$$j = x)}" + {(j,x). j\Basis \ (interval_lowerbound k)\j < x \ x < (interval_upperbound k)\j \ + (\i\d. (interval_lowerbound i)\j = x \ (interval_upperbound i)\j = x)}" lemma division_points_finite: fixes i::"('a::ordered_euclidean_space) set" assumes "d division_of i" shows "finite (division_points i d)" proof- note assm = division_ofD[OF assms] - let ?M = "\j. {(j,x)|x. (interval_lowerbound i)$$j < x \ x < (interval_upperbound i)$$j \ - (\i\d. (interval_lowerbound i)$$j = x \ (interval_upperbound i)$$j = x)}" - have *:"division_points i d = \(?M ` {..(?M ` Basis)" unfolding division_points_def by auto show ?thesis unfolding * using assm by auto qed lemma division_points_subset: fixes a::"'a::ordered_euclidean_space" - assumes "d division_of {a..b}" "\i {x. x$$k \ c}) {l \ {x. x$$k \ c} | l . l \ d \ ~(l \ {x. x$$k \ c} = {})} + assumes "d division_of {a..b}" "\i\Basis. a\i < b\i" "a\k < c" "c < b\k" and k:"k\Basis" + shows "division_points ({a..b} \ {x. x\k \ c}) {l \ {x. x\k \ c} | l . l \ d \ ~(l \ {x. x\k \ c} = {})} \ division_points ({a..b}) d" (is ?t1) and - "division_points ({a..b} \ {x. x$$k \ c}) {l \ {x. x$$k \ c} | l . l \ d \ ~(l \ {x. x$$k \ c} = {})} + "division_points ({a..b} \ {x. x\k \ c}) {l \ {x. x\k \ c} | l . l \ d \ ~(l \ {x. x\k \ c} = {})} \ division_points ({a..b}) d" (is ?t2) proof- note assm = division_ofD[OF assms(1)] - have *:"\i b$$i" "\i ((\\ i. if i = k then min (b $$ k) c else b $$ i)::'a) $$ i" - "\i\ i. if i = k then max (a $$ k) c else a $$ i)::'a) $$ i \ b$$i" "min (b $$ k) c = c" "max (a $$ k) c = c" + have *:"\i\Basis. a\i \ b\i" + "\i\Basis. a\i \ (\i\Basis. (if i = k then min (b \ k) c else b \ i) *\<^sub>R i) \ i" + "\i\Basis. (\i\Basis. (if i = k then max (a \ k) c else a \ i) *\<^sub>R i) \ i \ b\i" + "min (b \ k) c = c" "max (a \ k) c = c" using assms using less_imp_le by auto - show ?t1 unfolding division_points_def interval_split[OF k, of a b] - unfolding interval_bounds[OF *(1)] interval_bounds[OF *(2)] interval_bounds[OF *(3)] unfolding * - unfolding subset_eq apply(rule) unfolding mem_Collect_eq split_beta apply(erule bexE conjE)+ - unfolding mem_Collect_eq apply(erule exE conjE)+ unfolding euclidean_lambda_beta' - proof- fix i l x assume as:"a $$ fst x < snd x" "snd x < (if fst x = k then c else b $$ fst x)" - "interval_lowerbound i $$ fst x = snd x \ interval_upperbound i $$ fst x = snd x" - "i = l \ {x. x $$ k \ c}" "l \ d" "l \ {x. x $$ k \ c} \ {}" and fstx:"fst x fst x < snd x" "snd x < (if fst x = k then c else b \ fst x)" + "interval_lowerbound i \ fst x = snd x \ interval_upperbound i \ fst x = snd x" + "i = l \ {x. x \ k \ c}" "l \ d" "l \ {x. x \ k \ c} \ {}" and fstx:"fst x \Basis" from assm(4)[OF this(5)] guess u v apply-by(erule exE)+ note l=this - have *:"\i ((\\ i. if i = k then min (v $$ k) c else v $$ i)::'a) $$ i" + have *:"\i\Basis. u \ i \ (\i\Basis. (if i = k then min (v \ k) c else v \ i) *\<^sub>R i) \ i" using as(6) unfolding l interval_split[OF k] interval_ne_empty as . - have **:"\i v$$i" using l using as(6) unfolding interval_ne_empty[THEN sym] by auto - show "fst x a $$ fst x < snd x \ snd x < b $$ fst x \ (\i\d. interval_lowerbound i $$ fst x = snd x - \ interval_upperbound i $$ fst x = snd x)" apply(rule,rule fstx) - using as(1-3,5) unfolding l interval_split[OF k] interval_ne_empty as interval_bounds[OF *] apply- - apply(rule,assumption,rule) defer apply(rule_tac x="{u..v}" in bexI) unfolding interval_bounds[OF **] - apply(case_tac[!] "fst x = k") using assms fstx apply- unfolding euclidean_lambda_beta by auto + have **:"\i\Basis. u\i \ v\i" using l using as(6) unfolding interval_ne_empty[THEN sym] by auto + show "\i\d. interval_lowerbound i \ fst x = snd x \ interval_upperbound i \ fst x = snd x" + apply (rule bexI[OF _ `l \ d`]) + using as(1-3,5) fstx + unfolding l interval_bounds[OF **] interval_bounds[OF *] interval_split[OF k] as + by (auto split: split_if_asm) + show "snd x < b \ fst x" + using as(2) `c < b\k` by (auto split: split_if_asm) qed - show ?t2 unfolding division_points_def interval_split[OF k, of a b] + show ?t2 + unfolding division_points_def interval_split[OF k, of a b] unfolding interval_bounds[OF *(1)] interval_bounds[OF *(2)] interval_bounds[OF *(3)] unfolding * - unfolding subset_eq apply(rule) unfolding mem_Collect_eq split_beta apply(erule bexE conjE)+ - unfolding mem_Collect_eq apply(erule exE conjE)+ unfolding euclidean_lambda_beta' apply(rule,assumption) - proof- fix i l x assume as:"(if fst x = k then c else a $$ fst x) < snd x" "snd x < b $$ fst x" - "interval_lowerbound i $$ fst x = snd x \ interval_upperbound i $$ fst x = snd x" - "i = l \ {x. c \ x $$ k}" "l \ d" "l \ {x. c \ x $$ k} \ {}" and fstx:"fst x < DIM('a)" + unfolding subset_eq apply(rule) unfolding mem_Collect_eq split_beta + apply(erule bexE conjE)+ + apply(simp only: mem_Collect_eq inner_setsum_left_Basis simp_thms) + apply(erule exE conjE)+ + proof + fix i l x assume as:"(if fst x = k then c else a \ fst x) < snd x" "snd x < b \ fst x" + "interval_lowerbound i \ fst x = snd x \ interval_upperbound i \ fst x = snd x" + "i = l \ {x. c \ x \ k}" "l \ d" "l \ {x. c \ x \ k} \ {}" and fstx:"fst x \ Basis" from assm(4)[OF this(5)] guess u v apply-by(erule exE)+ note l=this - have *:"\i\ i. if i = k then max (u $$ k) c else u $$ i)::'a) $$ i \ v $$ i" + have *:"\i\Basis. (\i\Basis. (if i = k then max (u \ k) c else u \ i) *\<^sub>R i) \ i \ v \ i" using as(6) unfolding l interval_split[OF k] interval_ne_empty as . - have **:"\i v$$i" using l using as(6) unfolding interval_ne_empty[THEN sym] by auto - show "a $$ fst x < snd x \ snd x < b $$ fst x \ (\i\d. interval_lowerbound i $$ fst x = snd x \ - interval_upperbound i $$ fst x = snd x)" - using as(1-3,5) unfolding l interval_split[OF k] interval_ne_empty as interval_bounds[OF *] apply- - apply rule defer apply(rule,assumption) apply(rule_tac x="{u..v}" in bexI) unfolding interval_bounds[OF **] - apply(case_tac[!] "fst x = k") using assms fstx apply- by(auto simp add:euclidean_lambda_beta'[OF k]) qed qed + have **:"\i\Basis. u\i \ v\i" using l using as(6) unfolding interval_ne_empty[THEN sym] by auto + show "\i\d. interval_lowerbound i \ fst x = snd x \ interval_upperbound i \ fst x = snd x" + apply (rule bexI[OF _ `l \ d`]) + using as(1-3,5) fstx + unfolding l interval_bounds[OF **] interval_bounds[OF *] interval_split[OF k] as + by (auto split: split_if_asm) + show "a \ fst x < snd x" + using as(1) `a\k < c` by (auto split: split_if_asm) + qed +qed lemma division_points_psubset: fixes a::"'a::ordered_euclidean_space" - assumes "d division_of {a..b}" "\i d" "interval_lowerbound l$$k = c \ interval_upperbound l$$k = c" and k:"k {x. x$$k \ c}) {l \ {x. x$$k \ c} | l. l\d \ l \ {x. x$$k \ c} \ {}} + assumes "d division_of {a..b}" "\i\Basis. a\i < b\i" "a\k < c" "c < b\k" + "l \ d" "interval_lowerbound l\k = c \ interval_upperbound l\k = c" and k:"k\Basis" + shows "division_points ({a..b} \ {x. x\k \ c}) {l \ {x. x\k \ c} | l. l\d \ l \ {x. x\k \ c} \ {}} \ division_points ({a..b}) d" (is "?D1 \ ?D") - "division_points ({a..b} \ {x. x$$k \ c}) {l \ {x. x$$k \ c} | l. l\d \ l \ {x. x$$k \ c} \ {}} + "division_points ({a..b} \ {x. x\k \ c}) {l \ {x. x\k \ c} | l. l\d \ l \ {x. x\k \ c} \ {}} \ division_points ({a..b}) d" (is "?D2 \ ?D") -proof- have ab:"\i b$$i" using assms(2) by(auto intro!:less_imp_le) +proof- have ab:"\i\Basis. a\i \ b\i" using assms(2) by(auto intro!:less_imp_le) guess u v using division_ofD(4)[OF assms(1,5)] apply-by(erule exE)+ note l=this - have uv:"\i v$$i" "\i u$$i \ v$$i \ b$$i" + have uv:"\i\Basis. u\i \ v\i" "\i\Basis. a\i \ u\i \ v\i \ b\i" using division_ofD(2,2,3)[OF assms(1,5)] unfolding l interval_ne_empty unfolding subset_eq apply- defer apply(erule_tac x=u in ballE, erule_tac x=v in ballE) unfolding mem_interval by auto - have *:"interval_upperbound ({a..b} \ {x. x $$ k \ interval_upperbound l $$ k}) $$ k = interval_upperbound l $$ k" - "interval_upperbound ({a..b} \ {x. x $$ k \ interval_lowerbound l $$ k}) $$ k = interval_lowerbound l $$ k" + have *:"interval_upperbound ({a..b} \ {x. x \ k \ interval_upperbound l \ k}) \ k = interval_upperbound l \ k" + "interval_upperbound ({a..b} \ {x. x \ k \ interval_lowerbound l \ k}) \ k = interval_lowerbound l \ k" unfolding interval_split[OF k] apply(subst interval_bounds) prefer 3 apply(subst interval_bounds) unfolding l interval_bounds[OF uv(1)] using uv[rule_format,of k] ab k by auto have "\x. x \ ?D - ?D1" using assms(2-) apply-apply(erule disjE) - apply(rule_tac x="(k,(interval_lowerbound l)$$k)" in exI) defer - apply(rule_tac x="(k,(interval_upperbound l)$$k)" in exI) + apply(rule_tac x="(k,(interval_lowerbound l)\k)" in exI) defer + apply(rule_tac x="(k,(interval_upperbound l)\k)" in exI) unfolding division_points_def unfolding interval_bounds[OF ab] by(auto simp add:*) thus "?D1 \ ?D" apply-apply(rule,rule division_points_subset[OF assms(1-4)]) using k by auto - have *:"interval_lowerbound ({a..b} \ {x. x $$ k \ interval_lowerbound l $$ k}) $$ k = interval_lowerbound l $$ k" - "interval_lowerbound ({a..b} \ {x. x $$ k \ interval_upperbound l $$ k}) $$ k = interval_upperbound l $$ k" + have *:"interval_lowerbound ({a..b} \ {x. x \ k \ interval_lowerbound l \ k}) \ k = interval_lowerbound l \ k" + "interval_lowerbound ({a..b} \ {x. x \ k \ interval_upperbound l \ k}) \ k = interval_upperbound l \ k" unfolding interval_split[OF k] apply(subst interval_bounds) prefer 3 apply(subst interval_bounds) unfolding l interval_bounds[OF uv(1)] using uv[rule_format,of k] ab k by auto have "\x. x \ ?D - ?D2" using assms(2-) apply-apply(erule disjE) - apply(rule_tac x="(k,(interval_lowerbound l)$$k)" in exI) defer - apply(rule_tac x="(k,(interval_upperbound l)$$k)" in exI) + apply(rule_tac x="(k,(interval_lowerbound l)\k)" in exI) defer + apply(rule_tac x="(k,(interval_upperbound l)\k)" in exI) unfolding division_points_def unfolding interval_bounds[OF ab] by(auto simp add:*) thus "?D2 \ ?D" apply-apply(rule,rule division_points_subset[OF assms(1-4) k]) by auto qed @@ -2548,42 +2466,47 @@ using operativeD(1)[OF assms(2)] x by auto qed qed } assume "content {a..b} \ 0" note ab = this[unfolded content_lt_nz[THEN sym] content_pos_lt_eq] - hence ab':"\i b$$i" by (auto intro!: less_imp_le) show ?case + hence ab':"\i\Basis. a\i \ b\i" by (auto intro!: less_imp_le) show ?case proof(cases "division_points {a..b} d = {}") case True have d':"\i\d. \u v. i = {u..v} \ - (\j v$$j = a$$j \ u$$j = b$$j \ v$$j = b$$j \ u$$j = a$$j \ v$$j = b$$j)" + (\j\Basis. u\j = a\j \ v\j = a\j \ u\j = b\j \ v\j = b\j \ u\j = a\j \ v\j = b\j)" unfolding forall_in_division[OF goal1(4)] apply(rule,rule,rule) - apply(rule_tac x=a in exI,rule_tac x=b in exI) apply(rule,rule refl) apply(rule,rule) - proof- fix u v j assume j:"j d" note division_ofD(3)[OF goal1(4) this] - hence uv:"\i v$$i" "u$$j \ v$$j" using j unfolding interval_ne_empty by auto - have *:"\p r Q. \ j p \ r \ (\x\d. Q x) \ p \ r \ (Q {u..v})" using as j by auto - have "(j, u$$j) \ division_points {a..b} d" - "(j, v$$j) \ division_points {a..b} d" using True by auto + apply(rule_tac x=a in exI,rule_tac x=b in exI) apply(rule,rule refl) + proof + fix u v and j :: 'a assume j:"j\Basis" assume as:"{u..v} \ d" note division_ofD(3)[OF goal1(4) this] + hence uv:"\i\Basis. u\i \ v\i" "u\j \ v\j" using j unfolding interval_ne_empty by auto + have *:"\p r Q. \ j\Basis \ p \ r \ (\x\d. Q x) \ p \ r \ (Q {u..v})" using as j by auto + have "(j, u\j) \ division_points {a..b} d" + "(j, v\j) \ division_points {a..b} d" using True by auto note this[unfolded de_Morgan_conj division_points_def mem_Collect_eq split_conv interval_bounds[OF ab'] bex_simps] note *[OF this(1)] *[OF this(2)] note this[unfolded interval_bounds[OF uv(1)]] - moreover have "a$$j \ u$$j" "v$$j \ b$$j" using division_ofD(2,2,3)[OF goal1(4) as] + moreover have "a\j \ u\j" "v\j \ b\j" using division_ofD(2,2,3)[OF goal1(4) as] unfolding subset_eq apply- apply(erule_tac x=u in ballE,erule_tac[3] x=v in ballE) unfolding interval_ne_empty mem_interval using j by auto - ultimately show "u$$j = a$$j \ v$$j = a$$j \ u$$j = b$$j \ v$$j = b$$j \ u$$j = a$$j \ v$$j = b$$j" + ultimately show "u\j = a\j \ v\j = a\j \ u\j = b\j \ v\j = b\j \ u\j = a\j \ v\j = b\j" unfolding not_less de_Morgan_disj using ab[rule_format,of j] uv(2) j by auto - qed have "(1/2) *\<^sub>R (a+b) \ {a..b}" unfolding mem_interval using ab by(auto intro!:less_imp_le) + qed + have "(1/2) *\<^sub>R (a+b) \ {a..b}" + unfolding mem_interval using ab by(auto intro!: less_imp_le simp: inner_simps) note this[unfolded division_ofD(6)[OF goal1(4),THEN sym] Union_iff] then guess i .. note i=this guess u v using d'[rule_format,OF i(1)] apply-by(erule exE conjE)+ note uv=this have "{a..b} \ d" proof- { presume "i = {a..b}" thus ?thesis using i by auto } { presume "u = a" "v = b" thus "i = {a..b}" using uv by auto } - show "u = a" "v = b" unfolding euclidean_eq[where 'a='a] - proof(safe) fix j assume j:"jBasis" + note i(2)[unfolded uv mem_interval,rule_format,of j] + thus "u \ j = a \ j" "v \ j = b \ j" using uv(2)[rule_format,of j] j by (auto simp: inner_simps) qed qed hence *:"d = insert {a..b} (d - {{a..b}})" by auto have "iterate opp (d - {{a..b}}) f = neutral opp" apply(rule iterate_eq_neutral[OF goal1(2)]) proof fix x assume x:"x \ d - {{a..b}}" hence "x\d" by auto note d'[rule_format,OF this] then guess u v apply-by(erule exE conjE)+ note uv=this have "u\a \ v\b" using x[unfolded uv] by auto - then obtain j where "u$$j \ a$$j \ v$$j \ b$$j" and j:"j "{l \ {x. x$$k \ c} | l. l \ d \ l \ {x. x$$k \ c} \ {}}" - def d2 \ "{l \ {x. x$$k \ c} | l. l \ d \ l \ {x. x$$k \ c} \ {}}" - def cb \ "(\\ i. if i = k then c else b$$i)::'a" and ca \ "(\\ i. if i = k then c else a$$i)::'a" + def d1 \ "{l \ {x. x\k \ c} | l. l \ d \ l \ {x. x\k \ c} \ {}}" + def d2 \ "{l \ {x. x\k \ c} | l. l \ d \ l \ {x. x\k \ c} \ {}}" + def cb \ "(\i\Basis. (if i = k then c else b\i) *\<^sub>R i)::'a" + def ca \ "(\i\Basis. (if i = k then c else a\i) *\<^sub>R i)::'a" note division_points_psubset[OF goal1(4) ab kc(1-2) j] note psubset_card_mono[OF _ this(1)] psubset_card_mono[OF _ this(2)] - hence *:"(iterate opp d1 f) = f ({a..b} \ {x. x$$k \ c})" "(iterate opp d2 f) = f ({a..b} \ {x. x$$k \ c})" + hence *:"(iterate opp d1 f) = f ({a..b} \ {x. x\k \ c})" "(iterate opp d2 f) = f ({a..b} \ {x. x\k \ c})" apply- unfolding interval_split[OF kc(4)] apply(rule_tac[!] goal1(1)[rule_format]) using division_split[OF goal1(4), where k=k and c=c] unfolding interval_split[OF kc(4)] d1_def[symmetric] d2_def[symmetric] unfolding goal1(2) Suc_le_mono using goal1(2-3) using division_points_finite[OF goal1(4)] using kc(4) by auto have "f {a..b} = opp (iterate opp d1 f) (iterate opp d2 f)" (is "_ = ?prev") unfolding * apply(rule operativeD(2)) using goal1(3) using kc(4) by auto - also have "iterate opp d1 f = iterate opp d (\l. f(l \ {x. x$$k \ c}))" + also have "iterate opp d1 f = iterate opp d (\l. f(l \ {x. x\k \ c}))" unfolding d1_def apply(rule iterate_nonzero_image_lemma[unfolded o_def]) unfolding empty_as_interval apply(rule goal1 division_of_finite operativeD[OF goal1(3)])+ unfolding empty_as_interval[THEN sym] apply(rule content_empty) - proof(rule,rule,rule,erule conjE) fix l y assume as:"l \ d" "y \ d" "l \ {x. x $$ k \ c} = y \ {x. x $$ k \ c}" "l \ y" + proof(rule,rule,rule,erule conjE) fix l y assume as:"l \ d" "y \ d" "l \ {x. x \ k \ c} = y \ {x. x \ k \ c}" "l \ y" from division_ofD(4)[OF goal1(4) this(1)] guess u v apply-by(erule exE)+ note l=this - show "f (l \ {x. x $$ k \ c}) = neutral opp" unfolding l interval_split[OF kc(4)] + show "f (l \ {x. x \ k \ c}) = neutral opp" unfolding l interval_split[OF kc(4)] apply(rule operativeD(1) goal1)+ unfolding interval_split[THEN sym,OF kc(4)] apply(rule division_split_left_inj) apply(rule goal1) unfolding l[THEN sym] apply(rule as(1),rule as(2)) by(rule kc(4) as)+ - qed also have "iterate opp d2 f = iterate opp d (\l. f(l \ {x. x$$k \ c}))" + qed also have "iterate opp d2 f = iterate opp d (\l. f(l \ {x. x\k \ c}))" unfolding d2_def apply(rule iterate_nonzero_image_lemma[unfolded o_def]) unfolding empty_as_interval apply(rule goal1 division_of_finite operativeD[OF goal1(3)])+ unfolding empty_as_interval[THEN sym] apply(rule content_empty) - proof(rule,rule,rule,erule conjE) fix l y assume as:"l \ d" "y \ d" "l \ {x. c \ x $$ k} = y \ {x. c \ x $$ k}" "l \ y" + proof(rule,rule,rule,erule conjE) fix l y assume as:"l \ d" "y \ d" "l \ {x. c \ x \ k} = y \ {x. c \ x \ k}" "l \ y" from division_ofD(4)[OF goal1(4) this(1)] guess u v apply-by(erule exE)+ note l=this - show "f (l \ {x. x $$ k \ c}) = neutral opp" unfolding l interval_split[OF kc(4)] + show "f (l \ {x. x \ k \ c}) = neutral opp" unfolding l interval_split[OF kc(4)] apply(rule operativeD(1) goal1)+ unfolding interval_split[THEN sym,OF kc(4)] apply(rule division_split_right_inj) apply(rule goal1) unfolding l[THEN sym] apply(rule as(1),rule as(2)) by(rule as kc(4))+ - qed also have *:"\x\d. f x = opp (f (x \ {x. x $$ k \ c})) (f (x \ {x. c \ x $$ k}))" + qed also have *:"\x\d. f x = opp (f (x \ {x. x \ k \ c})) (f (x \ {x. c \ x \ k}))" unfolding forall_in_division[OF goal1(4)] apply(rule,rule,rule,rule operativeD(2)) using goal1(3) kc by auto - have "opp (iterate opp d (\l. f (l \ {x. x $$ k \ c}))) (iterate opp d (\l. f (l \ {x. c \ x $$ k}))) + have "opp (iterate opp d (\l. f (l \ {x. x \ k \ c}))) (iterate opp d (\l. f (l \ {x. c \ x \ k}))) = iterate opp d f" apply(subst(3) iterate_eq[OF _ *[rule_format]]) prefer 3 apply(rule iterate_op[THEN sym]) using goal1 by auto finally show ?thesis by auto @@ -2754,31 +2678,34 @@ subsection {* Similar theorems about relationship among components. *} lemma rsum_component_le: fixes f::"'a::ordered_euclidean_space \ 'b::euclidean_space" - assumes "p tagged_division_of {a..b}" "\x\{a..b}. (f x)$$i \ (g x)$$i" - shows "(setsum (\(x,k). content k *\<^sub>R f x) p)$$i \ (setsum (\(x,k). content k *\<^sub>R g x) p)$$i" - unfolding euclidean_component_setsum apply(rule setsum_mono) apply safe + assumes "p tagged_division_of {a..b}" "\x\{a..b}. (f x)\i \ (g x)\i" + shows "(setsum (\(x,k). content k *\<^sub>R f x) p)\i \ (setsum (\(x,k). content k *\<^sub>R g x) p)\i" + unfolding inner_setsum_left apply(rule setsum_mono) apply safe proof- fix a b assume ab:"(a,b) \ p" note assm = tagged_division_ofD(2-4)[OF assms(1) ab] from this(3) guess u v apply-by(erule exE)+ note b=this - show "(content b *\<^sub>R f a) $$ i \ (content b *\<^sub>R g a) $$ i" unfolding b - unfolding euclidean_simps real_scaleR_def apply(rule mult_left_mono) + show "(content b *\<^sub>R f a) \ i \ (content b *\<^sub>R g a) \ i" unfolding b + unfolding inner_simps real_scaleR_def apply(rule mult_left_mono) defer apply(rule content_pos_le,rule assms(2)[rule_format]) using assm by auto qed -lemma has_integral_component_le: fixes f g::"'a::ordered_euclidean_space \ 'b::euclidean_space" - assumes "(f has_integral i) s" "(g has_integral j) s" "\x\s. (f x)$$k \ (g x)$$k" - shows "i$$k \ j$$k" +lemma has_integral_component_le: + fixes f g::"'a::ordered_euclidean_space \ 'b::euclidean_space" + assumes k: "k \ Basis" + assumes "(f has_integral i) s" "(g has_integral j) s" "\x\s. (f x)\k \ (g x)\k" + shows "i\k \ j\k" proof - have lem:"\a b i (j::'b). \g f::'a \ 'b. (f has_integral i) ({a..b}) \ - (g has_integral j) ({a..b}) \ \x\{a..b}. (f x)$$k \ (g x)$$k \ i$$k \ j$$k" + (g has_integral j) ({a..b}) \ \x\{a..b}. (f x)\k \ (g x)\k \ i\k \ j\k" proof (rule ccontr) case goal1 - then have *: "0 < (i$$k - j$$k) / 3" by auto + then have *: "0 < (i\k - j\k) / 3" by auto guess d1 using goal1(1)[unfolded has_integral,rule_format,OF *] apply-by(erule exE conjE)+ note d1=this[rule_format] guess d2 using goal1(2)[unfolded has_integral,rule_format,OF *] apply-by(erule exE conjE)+ note d2=this[rule_format] guess p using fine_division_exists[OF gauge_inter[OF d1(1) d2(1)], of a b] unfolding fine_inter . - note p = this(1) conjunctD2[OF this(2)] note le_less_trans[OF component_le_norm, of _ _ k] term g + note p = this(1) conjunctD2[OF this(2)] + note le_less_trans[OF Basis_le_norm[OF k]] note this[OF d1(2)[OF conjI[OF p(1,2)]]] this[OF d2(2)[OF conjI[OF p(1,3)]]] thus False - unfolding euclidean_simps + unfolding inner_simps using rsum_component_le[OF p(1) goal1(3)] by (simp add: abs_real_def split: split_if_asm) qed let ?P = "\a b. s = {a..b}" @@ -2787,8 +2714,9 @@ show ?thesis apply(rule lem) using assms[unfolded s] by auto qed auto } assume as:"\ ?P" { presume "\ ?thesis \ False" thus ?thesis by auto } - assume "\ i$$k \ j$$k" hence ij:"(i$$k - j$$k) / 3 > 0" by auto - note has_integral_altD[OF _ as this] from this[OF assms(1)] this[OF assms(2)] guess B1 B2 . note B=this[rule_format] + assume "\ i\k \ j\k" hence ij:"(i\k - j\k) / 3 > 0" by auto + note has_integral_altD[OF _ as this] + from this[OF assms(2)] this[OF assms(3)] guess B1 B2 . note B=this[rule_format] have "bounded (ball 0 B1 \ ball (0::'a) B2)" unfolding bounded_Un by(rule conjI bounded_ball)+ from bounded_subset_closed_interval[OF this] guess a b apply- by(erule exE)+ note ab = conjunctD2[OF this[unfolded Un_subset_iff]] @@ -2796,69 +2724,50 @@ guess w2 using B(4)[OF ab(2)] .. note w2=conjunctD2[OF this] have *:"\w1 w2 j i::real .\w1 - i\ < (i - j) / 3 \ \w2 - j\ < (i - j) / 3 \ w1 \ w2 \ False" by (simp add: abs_real_def split: split_if_asm) - note le_less_trans[OF component_le_norm[of _ k]] note this[OF w1(2)] this[OF w2(2)] moreover - have "w1$$k \ w2$$k" apply(rule lem[OF w1(1) w2(1)]) using assms by auto ultimately - show False unfolding euclidean_simps by(rule *) qed + note le_less_trans[OF Basis_le_norm[OF k]] note this[OF w1(2)] this[OF w2(2)] moreover + have "w1\k \ w2\k" apply(rule lem[OF w1(1) w2(1)]) using assms by auto ultimately + show False unfolding inner_simps by(rule *) +qed lemma integral_component_le: fixes g f::"'a::ordered_euclidean_space \ 'b::euclidean_space" - assumes "f integrable_on s" "g integrable_on s" "\x\s. (f x)$$k \ (g x)$$k" - shows "(integral s f)$$k \ (integral s g)$$k" + assumes "k\Basis" "f integrable_on s" "g integrable_on s" "\x\s. (f x)\k \ (g x)\k" + shows "(integral s f)\k \ (integral s g)\k" apply(rule has_integral_component_le) using integrable_integral assms by auto -(*lemma has_integral_dest_vec1_le: fixes f::"'a::ordered_euclidean_space \ real^1" - assumes "(f has_integral i) s" "(g has_integral j) s" "\x\s. f x \ g x" - shows "dest_vec1 i \ dest_vec1 j" apply(rule has_integral_component_le[OF assms(1-2)]) - using assms(3) unfolding vector_le_def by auto - -lemma integral_dest_vec1_le: fixes f::"real^'n \ real^1" - assumes "f integrable_on s" "g integrable_on s" "\x\s. f x \ g x" - shows "dest_vec1(integral s f) \ dest_vec1(integral s g)" - apply(rule has_integral_dest_vec1_le) apply(rule_tac[1-2] integrable_integral) using assms by auto*) - lemma has_integral_component_nonneg: fixes f::"'a::ordered_euclidean_space \ 'b::euclidean_space" - assumes "(f has_integral i) s" "\x\s. 0 \ (f x)$$k" shows "0 \ i$$k" - using has_integral_component_le[OF has_integral_0 assms(1)] using assms(2-) by auto + assumes "k\Basis" "(f has_integral i) s" "\x\s. 0 \ (f x)\k" shows "0 \ i\k" + using has_integral_component_le[OF assms(1) has_integral_0 assms(2)] using assms(3-) by auto lemma integral_component_nonneg: fixes f::"'a::ordered_euclidean_space \ 'b::euclidean_space" - assumes "f integrable_on s" "\x\s. 0 \ (f x)$$k" shows "0 \ (integral s f)$$k" + assumes "k\Basis" "f integrable_on s" "\x\s. 0 \ (f x)\k" shows "0 \ (integral s f)\k" apply(rule has_integral_component_nonneg) using assms by auto -(*lemma has_integral_dest_vec1_nonneg: fixes f::"real^'n \ real^1" - assumes "(f has_integral i) s" "\x\s. 0 \ f x" shows "0 \ i" - using has_integral_component_nonneg[OF assms(1), of 1] - using assms(2) unfolding vector_le_def by auto - -lemma integral_dest_vec1_nonneg: fixes f::"real^'n \ real^1" - assumes "f integrable_on s" "\x\s. 0 \ f x" shows "0 \ integral s f" - apply(rule has_integral_dest_vec1_nonneg) using assms by auto*) - lemma has_integral_component_neg: fixes f::"'a::ordered_euclidean_space \ 'b::ordered_euclidean_space" - assumes "(f has_integral i) s" "\x\s. (f x)$$k \ 0"shows "i$$k \ 0" - using has_integral_component_le[OF assms(1) has_integral_0] assms(2-) by auto - -(*lemma has_integral_dest_vec1_neg: fixes f::"real^'n \ real^1" - assumes "(f has_integral i) s" "\x\s. f x \ 0" shows "i \ 0" - using has_integral_component_neg[OF assms(1),of 1] using assms(2) by auto*) - -lemma has_integral_component_lbound: fixes f::"'a::ordered_euclidean_space => 'b::ordered_euclidean_space" - assumes "(f has_integral i) {a..b}" "\x\{a..b}. B \ f(x)$$k" "k i$$k" - using has_integral_component_le[OF has_integral_const assms(1),of "(\\ i. B)::'b" k] assms(2-) - unfolding euclidean_simps euclidean_lambda_beta'[OF assms(3)] by(auto simp add:field_simps) - -lemma has_integral_component_ubound: fixes f::"'a::ordered_euclidean_space => 'b::ordered_euclidean_space" - assumes "(f has_integral i) {a..b}" "\x\{a..b}. f x$$k \ B" "k B * content({a..b})" - using has_integral_component_le[OF assms(1) has_integral_const, of k "\\ i. B"] - unfolding euclidean_simps euclidean_lambda_beta'[OF assms(3)] using assms(2) by(auto simp add:field_simps) + assumes "k\Basis" "(f has_integral i) s" "\x\s. (f x)\k \ 0"shows "i\k \ 0" + using has_integral_component_le[OF assms(1,2) has_integral_0] assms(2-) by auto + +lemma has_integral_component_lbound: + fixes f::"'a::ordered_euclidean_space => 'b::ordered_euclidean_space" + assumes "(f has_integral i) {a..b}" "\x\{a..b}. B \ f(x)\k" "k\Basis" + shows "B * content {a..b} \ i\k" + using has_integral_component_le[OF assms(3) has_integral_const assms(1),of "(\i\Basis. B *\<^sub>R i)::'b"] assms(2-) + by (auto simp add:field_simps) + +lemma has_integral_component_ubound: + fixes f::"'a::ordered_euclidean_space => 'b::ordered_euclidean_space" + assumes "(f has_integral i) {a..b}" "\x\{a..b}. f x\k \ B" "k\Basis" + shows "i\k \ B * content({a..b})" + using has_integral_component_le[OF assms(3,1) has_integral_const, of "\i\Basis. B *\<^sub>R i"] assms(2-) + by(auto simp add:field_simps) lemma integral_component_lbound: fixes f::"'a::ordered_euclidean_space => 'b::ordered_euclidean_space" - assumes "f integrable_on {a..b}" "\x\{a..b}. B \ f(x)$$k" "k (integral({a..b}) f)$$k" + assumes "f integrable_on {a..b}" "\x\{a..b}. B \ f(x)\k" "k\Basis" + shows "B * content({a..b}) \ (integral({a..b}) f)\k" apply(rule has_integral_component_lbound) using assms unfolding has_integral_integral by auto lemma integral_component_ubound: fixes f::"'a::ordered_euclidean_space => 'b::ordered_euclidean_space" - assumes "f integrable_on {a..b}" "\x\{a..b}. f(x)$$k \ B" "k B * content({a..b})" + assumes "f integrable_on {a..b}" "\x\{a..b}. f(x)\k \ B" "k\Basis" + shows "(integral({a..b}) f)\k \ B * content({a..b})" apply(rule has_integral_component_ubound) using assms unfolding has_integral_integral by auto subsection {* Uniform limit of integrable functions is integrable. *} @@ -2949,68 +2858,76 @@ apply(rule iterate_nonzero_image_lemma) apply(rule assms monoidal_monoid)+ unfolding assms using neutral_add unfolding neutral_add using assms by auto -lemma interval_doublesplit: fixes a::"'a::ordered_euclidean_space" assumes "k {x . abs(x$$k - c) \ (e::real)} = - {(\\ i. if i = k then max (a$$k) (c - e) else a$$i) .. (\\ i. if i = k then min (b$$k) (c + e) else b$$i)}" +lemma interval_doublesplit: fixes a::"'a::ordered_euclidean_space" assumes "k\Basis" + shows "{a..b} \ {x . abs(x\k - c) \ (e::real)} = + {(\i\Basis. (if i = k then max (a\k) (c - e) else a\i) *\<^sub>R i) .. + (\i\Basis. (if i = k then min (b\k) (c + e) else b\i) *\<^sub>R i)}" proof- have *:"\x c e::real. abs(x - c) \ e \ x \ c - e \ x \ c + e" by auto have **:"\s P Q. s \ {x. P x \ Q x} = (s \ {x. Q x}) \ {x. P x}" by blast show ?thesis unfolding * ** interval_split[OF assms] by(rule refl) qed -lemma division_doublesplit: fixes a::"'a::ordered_euclidean_space" assumes "p division_of {a..b}" and k:"k {x. abs(x$$k - c) \ e} |l. l \ p \ l \ {x. abs(x$$k - c) \ e} \ {}} division_of ({a..b} \ {x. abs(x$$k - c) \ e})" +lemma division_doublesplit: fixes a::"'a::ordered_euclidean_space" assumes "p division_of {a..b}" and k:"k\Basis" + shows "{l \ {x. abs(x\k - c) \ e} |l. l \ p \ l \ {x. abs(x\k - c) \ e} \ {}} division_of ({a..b} \ {x. abs(x\k - c) \ e})" proof- have *:"\x c. abs(x - c) \ e \ x \ c - e \ x \ c + e" by auto have **:"\p q p' q'. p division_of q \ p = p' \ q = q' \ p' division_of q'" by auto note division_split(1)[OF assms, where c="c+e",unfolded interval_split[OF k]] note division_split(2)[OF this, where c="c-e" and k=k,OF k] thus ?thesis apply(rule **) using k apply- unfolding interval_doublesplit unfolding * unfolding interval_split interval_doublesplit apply(rule set_eqI) unfolding mem_Collect_eq apply rule apply(erule conjE exE)+ apply(rule_tac x=la in exI) defer - apply(erule conjE exE)+ apply(rule_tac x="l \ {x. c + e \ x $$ k}" in exI) apply rule defer apply rule + apply(erule conjE exE)+ apply(rule_tac x="l \ {x. c + e \ x \ k}" in exI) apply rule defer apply rule apply(rule_tac x=l in exI) by blast+ qed -lemma content_doublesplit: fixes a::"'a::ordered_euclidean_space" assumes "0 < e" and k:"k {x. abs(x$$k - c) \ d}) < e" +lemma content_doublesplit: fixes a::"'a::ordered_euclidean_space" assumes "0 < e" and k:"k\Basis" + obtains d where "0 < d" "content({a..b} \ {x. abs(x\k - c) \ d}) < e" proof(cases "content {a..b} = 0") case True show ?thesis apply(rule that[of 1]) defer unfolding interval_doublesplit[OF k] apply(rule le_less_trans[OF content_subset]) defer apply(subst True) unfolding interval_doublesplit[THEN sym,OF k] using assms by auto -next case False def d \ "e / 3 / setprod (\i. b$$i - a$$i) ({.. "e / 3 / setprod (\i. b\i - a\i) (Basis - {k})" note False[unfolded content_eq_0 not_ex not_le, rule_format] - hence "\x. x b$$x > a$$x" by(auto simp add:not_le) - hence prod0:"0 < setprod (\i. b$$i - a$$i) ({..x. x\Basis \ b\x > a\x" by(auto simp add:not_le) + hence prod0:"0 < setprod (\i. b\i - a\i) (Basis - {k})" apply-apply(rule setprod_pos) by(auto simp add:field_simps) hence "d > 0" unfolding d_def using assms by(auto simp add:field_simps) thus ?thesis - proof(rule that[of d]) have *:"{.. {x. \x $$ k - c\ \ d} \ {} \ - (\i\{.. {x. \x $$ k - c\ \ d}) $$ i - - interval_lowerbound ({a..b} \ {x. \x $$ k - c\ \ d}) $$ i) - = (\i\{.. {x. \x \ k - c\ \ d} \ {} \ + (\i\Basis - {k}. interval_upperbound ({a..b} \ {x. \x \ k - c\ \ d}) \ i + - interval_lowerbound ({a..b} \ {x. \x \ k - c\ \ d}) \ i) + = (\i\Basis - {k}. b\i - a\i)" apply(rule setprod_cong,rule refl) unfolding interval_doublesplit[OF k] apply(subst interval_bounds) defer apply(subst interval_bounds) unfolding interval_eq_empty not_ex not_less by auto - show "content ({a..b} \ {x. \x $$ k - c\ \ d}) < e" apply(cases) unfolding content_def apply(subst if_P,assumption,rule assms) + show "content ({a..b} \ {x. \x \ k - c\ \ d}) < e" apply(cases) unfolding content_def apply(subst if_P,assumption,rule assms) unfolding if_not_P apply(subst *) apply(subst setprod_insert) unfolding ** unfolding interval_doublesplit[OF k] interval_eq_empty not_ex not_less prefer 3 - apply(subst interval_bounds) defer apply(subst interval_bounds) unfolding euclidean_lambda_beta'[OF k] if_P[OF refl] - proof- have "(min (b $$ k) (c + d) - max (a $$ k) (c - d)) \ 2 * d" by auto - also have "... < e / (\i\{..i\{.. k) (c + d) - max (a \ k) (c - d)) \ 2 * d" by auto + also have "... < e / (\i\Basis - {k}. b \ i - a \ i)" unfolding d_def using assms prod0 by(auto simp add:field_simps) + finally show "(min (b \ k) (c + d) - max (a \ k) (c - d)) * (\i\Basis - {k}. b \ i - a \ i) < e" + unfolding pos_less_divide_eq[OF prod0] . + qed auto + qed +qed + +lemma negligible_standard_hyperplane[intro]: + fixes k :: "'a::ordered_euclidean_space" + assumes k: "k \ Basis" + shows "negligible {x. x\k = c}" unfolding negligible_def has_integral apply(rule,rule,rule,rule) proof- case goal1 from content_doublesplit[OF this k,of a b c] guess d . note d=this - let ?i = "indicator {x::'a. x$$k = c} :: 'a\real" + let ?i = "indicator {x::'a. x\k = c} :: 'a\real" show ?case apply(rule_tac x="\x. ball x d" in exI) apply(rule,rule gauge_ball,rule d) proof(rule,rule) fix p assume p:"p tagged_division_of {a..b} \ (\x. ball x d) fine p" - have *:"(\(x, ka)\p. content ka *\<^sub>R ?i x) = (\(x, ka)\p. content (ka \ {x. abs(x$$k - c) \ d}) *\<^sub>R ?i x)" + have *:"(\(x, ka)\p. content ka *\<^sub>R ?i x) = (\(x, ka)\p. content (ka \ {x. abs(x\k - c) \ d}) *\<^sub>R ?i x)" apply(rule setsum_cong2) unfolding split_paired_all real_scaleR_def mult_cancel_right split_conv apply(cases,rule disjI1,assumption,rule disjI2) - proof- fix x l assume as:"(x,l)\p" "?i x \ 0" hence xk:"x$$k = c" unfolding indicator_def apply-by(rule ccontr,auto) - show "content l = content (l \ {x. \x $$ k - c\ \ d})" apply(rule arg_cong[where f=content]) + proof- fix x l assume as:"(x,l)\p" "?i x \ 0" hence xk:"x\k = c" unfolding indicator_def apply-by(rule ccontr,auto) + show "content l = content (l \ {x. \x \ k - c\ \ d})" apply(rule arg_cong[where f=content]) apply(rule set_eqI,rule,rule) unfolding mem_Collect_eq proof- fix y assume y:"y\l" note p[THEN conjunct2,unfolded fine_def,rule_format,OF as(1),unfolded split_conv] - note this[unfolded subset_eq mem_ball dist_norm,rule_format,OF y] note le_less_trans[OF component_le_norm[of _ k] this] - thus "\y $$ k - c\ \ d" unfolding euclidean_simps xk by auto + note this[unfolded subset_eq mem_ball dist_norm,rule_format,OF y] note le_less_trans[OF Basis_le_norm[OF k] this] + thus "\y \ k - c\ \ d" unfolding inner_simps xk by auto qed auto qed note p'= tagged_division_ofD[OF p[THEN conjunct1]] and p''=division_of_tagged_division[OF p[THEN conjunct1]] show "norm ((\(x, ka)\p. content ka *\<^sub>R ?i x) - 0) < e" unfolding diff_0_right * unfolding real_scaleR_def real_norm_def @@ -3018,33 +2935,33 @@ apply(rule mult_nonneg_nonneg) apply(drule p'(4)) apply(erule exE)+ apply(rule_tac b=b in back_subst) prefer 2 apply(subst(asm) eq_commute) apply assumption apply(subst interval_doublesplit[OF k]) apply(rule content_pos_le) apply(rule indicator_pos_le) - proof- have "(\(x, ka)\p. content (ka \ {x. \x $$ k - c\ \ d}) * ?i x) \ (\(x, ka)\p. content (ka \ {x. \x $$ k - c\ \ d}))" + proof- have "(\(x, ka)\p. content (ka \ {x. \x \ k - c\ \ d}) * ?i x) \ (\(x, ka)\p. content (ka \ {x. \x \ k - c\ \ d}))" apply(rule setsum_mono) unfolding split_paired_all split_conv apply(rule mult_right_le_one_le) apply(drule p'(4)) by(auto simp add:interval_doublesplit[OF k]) also have "... < e" apply(subst setsum_over_tagged_division_lemma[OF p[THEN conjunct1]]) - proof- case goal1 have "content ({u..v} \ {x. \x $$ k - c\ \ d}) \ content {u..v}" + proof- case goal1 have "content ({u..v} \ {x. \x \ k - c\ \ d}) \ content {u..v}" unfolding interval_doublesplit[OF k] apply(rule content_subset) unfolding interval_doublesplit[THEN sym,OF k] by auto thus ?case unfolding goal1 unfolding interval_doublesplit[OF k] by (blast intro: antisym) - next have *:"setsum content {l \ {x. \x $$ k - c\ \ d} |l. l \ snd ` p \ l \ {x. \x $$ k - c\ \ d} \ {}} \ 0" + next have *:"setsum content {l \ {x. \x \ k - c\ \ d} |l. l \ snd ` p \ l \ {x. \x \ k - c\ \ d} \ {}} \ 0" apply(rule setsum_nonneg,rule) unfolding mem_Collect_eq image_iff apply(erule exE bexE conjE)+ unfolding split_paired_all - proof- fix x l a b assume as:"x = l \ {x. \x $$ k - c\ \ d}" "(a, b) \ p" "l = snd (a, b)" + proof- fix x l a b assume as:"x = l \ {x. \x \ k - c\ \ d}" "(a, b) \ p" "l = snd (a, b)" guess u v using p'(4)[OF as(2)] apply-by(erule exE)+ note * = this show "content x \ 0" unfolding as snd_conv * interval_doublesplit[OF k] by(rule content_pos_le) qed have **:"norm (1::real) \ 1" by auto note division_doublesplit[OF p'' k,unfolded interval_doublesplit[OF k]] note dsum_bound[OF this **,unfolded interval_doublesplit[THEN sym,OF k]] note this[unfolded real_scaleR_def real_norm_def mult_1_right mult_1, of c d] note le_less_trans[OF this d(2)] - from this[unfolded abs_of_nonneg[OF *]] show "(\ka\snd ` p. content (ka \ {x. \x $$ k - c\ \ d})) < e" + from this[unfolded abs_of_nonneg[OF *]] show "(\ka\snd ` p. content (ka \ {x. \x \ k - c\ \ d})) < e" apply(subst vsum_nonzero_image_lemma[of "snd ` p" content "{}", unfolded o_def,THEN sym]) apply(rule finite_imageI p' content_empty)+ unfolding forall_in_division[OF p''] proof(rule,rule,rule,rule,rule,rule,rule,erule conjE) fix m n u v - assume as:"{m..n} \ snd ` p" "{u..v} \ snd ` p" "{m..n} \ {u..v}" "{m..n} \ {x. \x $$ k - c\ \ d} = {u..v} \ {x. \x $$ k - c\ \ d}" - have "({m..n} \ {x. \x $$ k - c\ \ d}) \ ({u..v} \ {x. \x $$ k - c\ \ d}) \ {m..n} \ {u..v}" by blast + assume as:"{m..n} \ snd ` p" "{u..v} \ snd ` p" "{m..n} \ {u..v}" "{m..n} \ {x. \x \ k - c\ \ d} = {u..v} \ {x. \x \ k - c\ \ d}" + have "({m..n} \ {x. \x \ k - c\ \ d}) \ ({u..v} \ {x. \x \ k - c\ \ d}) \ {m..n} \ {u..v}" by blast note interior_mono[OF this, unfolded division_ofD(5)[OF p'' as(1-3)] interior_inter[of "{m..n}"]] - hence "interior ({m..n} \ {x. \x $$ k - c\ \ d}) = {}" unfolding as Int_absorb by auto - thus "content ({m..n} \ {x. \x $$ k - c\ \ d}) = 0" unfolding interval_doublesplit[OF k] content_eq_0_interior[THEN sym] . + hence "interior ({m..n} \ {x. \x \ k - c\ \ d}) = {}" unfolding as Int_absorb by auto + thus "content ({m..n} \ {x. \x \ k - c\ \ d}) = 0" unfolding interval_doublesplit[OF k] content_eq_0_interior[THEN sym] . qed qed - finally show "(\(x, ka)\p. content (ka \ {x. \x $$ k - c\ \ d}) * ?i x) < e" . + finally show "(\(x, ka)\p. content (ka \ {x. \x \ k - c\ \ d}) * ?i x) < e" . qed qed qed subsection {* A technical lemma about "refinement" of division. *} @@ -3224,7 +3141,7 @@ using negligible_union by auto lemma negligible_sing[intro]: "negligible {a::_::ordered_euclidean_space}" - using negligible_standard_hyperplane[of 0 "a$$0"] by auto + using negligible_standard_hyperplane[OF SOME_Basis, of "a \ (SOME i. i \ Basis)"] by auto lemma negligible_insert[simp]: "negligible(insert a s) \ negligible s" apply(subst insert_is_Un) unfolding negligible_union_eq by auto @@ -3276,11 +3193,20 @@ subsection {* In particular, the boundary of an interval is negligible. *} lemma negligible_frontier_interval: "negligible({a::'a::ordered_euclidean_space..b} - {a<.. ?A" apply rule unfolding Diff_iff mem_interval not_all - apply(erule conjE exE)+ apply(rule_tac X="{x. x $$ xa = a $$ xa} \ {x. x $$ xa = b $$ xa}" in UnionI) - apply(erule_tac[!] x=xa in allE) by auto - thus ?thesis apply-apply(rule negligible_subset[of ?A]) apply(rule negligible_unions[OF finite_imageI]) by auto qed +proof- + let ?A = "\((\k. {x. x\k = a\k} \ {x::'a. x\k = b\k}) ` Basis)" + have "{a..b} - {a<.. ?A" + apply rule unfolding Diff_iff mem_interval + apply simp + apply(erule conjE bexE)+ + apply(rule_tac x=i in bexI) + by auto + thus ?thesis + apply- + apply(rule negligible_subset[of ?A]) + apply(rule negligible_unions[OF finite_imageI]) + by auto +qed lemma has_integral_spike_interior: assumes "\x\{a<.. e" fixes f::"'b::ordered_euclidean_space \ 'a::banach" shows "operative op \ (\i. \g. (\x\i. norm (f x - g (x::'b)) \ e) \ g integrable_on i)" unfolding operative_def neutral_and -proof safe fix a b::"'b" { assume "content {a..b} = 0" +proof safe + fix a b::"'b" + { assume "content {a..b} = 0" thus "\g. (\x\{a..b}. norm (f x - g x) \ e) \ g integrable_on {a..b}" apply(rule_tac x=f in exI) using assms by(auto intro!:integrable_on_null) } - { fix c k g assume as:"\x\{a..b}. norm (f x - g x) \ e" "g integrable_on {a..b}" and k:"kg. (\x\{a..b} \ {x. x $$ k \ c}. norm (f x - g x) \ e) \ g integrable_on {a..b} \ {x. x $$ k \ c}" - "\g. (\x\{a..b} \ {x. c \ x $$ k}. norm (f x - g x) \ e) \ g integrable_on {a..b} \ {x. c \ x $$ k}" + { fix c g and k :: 'b + assume as:"\x\{a..b}. norm (f x - g x) \ e" "g integrable_on {a..b}" and k:"k\Basis" + show "\g. (\x\{a..b} \ {x. x \ k \ c}. norm (f x - g x) \ e) \ g integrable_on {a..b} \ {x. x \ k \ c}" + "\g. (\x\{a..b} \ {x. c \ x \ k}. norm (f x - g x) \ e) \ g integrable_on {a..b} \ {x. c \ x \ k}" apply(rule_tac[!] x=g in exI) using as(1) integrable_split[OF as(2) k] by auto } - fix c k g1 g2 assume as:"\x\{a..b} \ {x. x $$ k \ c}. norm (f x - g1 x) \ e" "g1 integrable_on {a..b} \ {x. x $$ k \ c}" - "\x\{a..b} \ {x. c \ x $$ k}. norm (f x - g2 x) \ e" "g2 integrable_on {a..b} \ {x. c \ x $$ k}" - assume k:"kx\{a..b} \ {x. x \ k \ c}. norm (f x - g1 x) \ e" "g1 integrable_on {a..b} \ {x. x \ k \ c}" + "\x\{a..b} \ {x. c \ x \ k}. norm (f x - g2 x) \ e" "g2 integrable_on {a..b} \ {x. c \ x \ k}" + assume k:"k\Basis" + let ?g = "\x. if x\k = c then f x else if x\k \ c then g1 x else g2 x" show "\g. (\x\{a..b}. norm (f x - g x) \ e) \ g integrable_on {a..b}" apply(rule_tac x="?g" in exI) - proof safe case goal1 thus ?case apply- apply(cases "x$$k=c", case_tac "x$$k < c") using as assms by auto - next case goal2 presume "?g integrable_on {a..b} \ {x. x $$ k \ c}" "?g integrable_on {a..b} \ {x. x $$ k \ c}" + proof safe case goal1 thus ?case apply- apply(cases "x\k=c", case_tac "x\k < c") using as assms by auto + next case goal2 presume "?g integrable_on {a..b} \ {x. x \ k \ c}" "?g integrable_on {a..b} \ {x. x \ k \ c}" then guess h1 h2 unfolding integrable_on_def by auto from has_integral_split[OF this k] show ?case unfolding integrable_on_def by auto - next show "?g integrable_on {a..b} \ {x. x $$ k \ c}" "?g integrable_on {a..b} \ {x. x $$ k \ c}" + next show "?g integrable_on {a..b} \ {x. x \ k \ c}" "?g integrable_on {a..b} \ {x. x \ k \ c}" apply(rule_tac[!] integrable_spike[OF negligible_standard_hyperplane[of k c]]) using k as(2,4) by auto qed qed lemma approximable_on_division: fixes f::"'b::ordered_euclidean_space \ 'a::banach" @@ -3355,11 +3284,15 @@ subsection {* Specialization of additivity to one dimension. *} +lemma + shows real_inner_1_left: "inner 1 x = x" + and real_inner_1_right: "inner x 1 = x" + by simp_all + lemma operative_1_lt: assumes "monoidal opp" shows "operative opp f \ ((\a b. b \ a \ f {a..b::real} = neutral opp) \ (\a b c. a < c \ c < b \ opp (f{a..c})(f{c..b}) = f {a..b}))" - unfolding operative_def content_eq_0 DIM_real less_one simp_thms(39,41) Eucl_real_simps - (* The dnf_simps simplify "\ x. x= _ \ _" and "\k. k = _ \ _" *) + apply (simp add: operative_def content_eq_0 less_one) proof safe fix a b c::"real" assume as:"\a b c. f {a..b} = opp (f ({a..b} \ {x. x \ c})) (f ({a..b} \ {x. c \ x}))" "a < c" "c < b" from this(2-) have "{a..b} \ {x. x \ c} = {a..c}" "{a..b} \ {x. x \ c} = {c..b}" by auto @@ -3376,10 +3309,11 @@ show ?thesis unfolding * apply(subst as(1)[rule_format,of 0 1]) using assms by auto qed next case True hence *:"min (b) c = c" "max a c = c" by auto - have **:"0 < DIM(real)" by auto - have ***:"\P Q. (\\ i. if i = 0 then P i else Q i) = (P 0::real)" apply(subst euclidean_eq) - apply safe unfolding euclidean_lambda_beta' by auto - show ?thesis unfolding interval_split[OF **,unfolded Eucl_real_simps(1,3)] unfolding *** * + have **: "(1::real) \ Basis" by simp + have ***:"\P Q. (\i\Basis. (if i = 1 then P i else Q i) *\<^sub>R i) = (P 1::real)" + by simp + show ?thesis + unfolding interval_split[OF **, unfolded real_inner_1_right] unfolding *** * proof(cases "c = a \ c = b") case False thus "f {a..b} = opp (f {a..c}) (f {c..b})" apply-apply(subst as(2)[rule_format]) using True by auto @@ -3411,13 +3345,13 @@ subsection {* Special case of additivity we need for the FCT. *} lemma interval_bound_sing[simp]: "interval_upperbound {a} = a" "interval_lowerbound {a} = a" - unfolding interval_upperbound_def interval_lowerbound_def by auto + unfolding interval_upperbound_def interval_lowerbound_def by (auto simp: euclidean_representation) lemma additive_tagged_division_1: fixes f::"real \ 'a::real_normed_vector" assumes "a \ b" "p tagged_division_of {a..b}" shows "setsum (\(x,k). f(interval_upperbound k) - f(interval_lowerbound k)) p = f b - f a" proof- let ?f = "(\k::(real) set. if k = {} then 0 else f(interval_upperbound k) - f(interval_lowerbound k))" - have ***:"\i b $$ i" using assms by auto + have ***:"\i\Basis. a \ i \ b \ i" using assms by auto have *:"operative op + ?f" unfolding operative_1_lt[OF monoidal_monoid] interval_eq_empty by auto have **:"{a..b} \ {}" using assms(1) by auto note operative_tagged_division[OF monoidal_monoid * assms(2)] note * = this[unfolded if_not_P[OF **] interval_bounds[OF ***],THEN sym] @@ -3518,38 +3452,30 @@ have "k \ \(s - {k})" apply safe apply(rule *[unfolded closed_limpt,rule_format]) unfolding islimpt_approachable proof safe fix x and e::real assume as:"x\k" "e>0" from k(2)[unfolded k content_eq_0] guess i .. - hence i:"c$$i = d$$i" "i "(\\ j. if j = i then if c$$i \ (a$$i + b$$i) / 2 then c$$i + - min e (b$$i - c$$i) / 2 else c$$i - min e (c$$i - a$$i) / 2 else x$$j)::'a" + hence i:"c\i = d\i" "i\Basis" using s(3)[OF k(1),unfolded k] unfolding interval_ne_empty by auto + hence xi:"x\i = d\i" using as unfolding k mem_interval by (metis antisym) + def y \ "(\j\Basis. (if j = i then if c\i \ (a\i + b\i) / 2 then c\i + + min e (b\i - c\i) / 2 else c\i - min e (c\i - a\i) / 2 else x\j) *\<^sub>R j)::'a" show "\x'\\(s - {k}). x' \ x \ dist x' x < e" apply(rule_tac x=y in bexI) proof have "d \ {c..d}" using s(3)[OF k(1)] unfolding k interval_eq_empty mem_interval by(fastforce simp add: not_less) - hence "d \ {a..b}" using s(2)[OF k(1)] unfolding k by auto note di = this[unfolded mem_interval,THEN spec[where x=i]] - hence xyi:"y$$i \ x$$i" unfolding y_def unfolding i xi euclidean_lambda_beta'[OF i(2)] if_P[OF refl] - apply(cases) apply(subst if_P,assumption) unfolding if_not_P not_le using as(2) - using assms(2)[unfolded content_eq_0] using i(2) - by (auto simp add: not_le min_def) - thus "y \ x" unfolding euclidean_eq[where 'a='a] using i by auto - have *:"{..i. 0) {.. {a..b}" using s(2)[OF k(1)] unfolding k by auto note di = this[unfolded mem_interval,THEN bspec[where x=i]] + hence xyi:"y\i \ x\i" + unfolding y_def i xi using as(2) assms(2)[unfolded content_eq_0] i(2) + by (auto elim!: ballE[of _ _ i]) + thus "y \ x" unfolding euclidean_eq_iff[where 'a='a] using i by auto + have *:"Basis = insert i (Basis - {i})" using i by auto + have "norm (y - x) < e + setsum (\i. 0) Basis" apply(rule le_less_trans[OF norm_le_l1]) apply(subst *,subst setsum_insert) prefer 3 apply(rule add_less_le_mono) - proof- show "\(y - x) $$ i\ < e" unfolding y_def euclidean_simps euclidean_lambda_beta'[OF i(2)] if_P[OF refl] - apply(cases) apply(subst if_P,assumption) unfolding if_not_P unfolding i xi using di as(2) by auto - show "(\i\{..(y - x) $$ i\) \ (\i\{..(y - x) \ i\ < e" + using di as(2) y_def i xi by (auto simp: inner_simps) + show "(\i\Basis - {i}. \(y - x) \ i\) \ (\i\Basis. 0)" + unfolding y_def by (auto simp: inner_simps) qed auto thus "dist y x < e" unfolding dist_norm by auto - have "y\k" unfolding k mem_interval apply rule apply(erule_tac x=i in allE) using xyi unfolding k i xi by auto - moreover have "y \ \s" unfolding s mem_interval - proof(rule,rule) note simps = y_def euclidean_lambda_beta' if_not_P - fix j assume j:"j y $$ j \ y $$ j \ b $$ j" - proof(cases "j = i") case False have "x \ {a..b}" using s(2)[OF k(1)] as(1) by auto - thus ?thesis using j apply- unfolding simps if_not_P[OF False] unfolding mem_interval by auto - next case True note T = this show ?thesis - proof(cases "c $$ i \ (a $$ i + b $$ i) / 2") - case True show ?thesis unfolding simps if_P[OF T] if_P[OF True] unfolding i - using True as(2) di apply-apply rule unfolding T by (auto simp add:field_simps) - next case False thus ?thesis unfolding simps if_P[OF T] if_not_P[OF False] unfolding i - using True as(2) di apply-apply rule unfolding T by (auto simp add:field_simps) - qed qed qed + have "y\k" unfolding k mem_interval apply rule apply(erule_tac x=i in ballE) using xyi k i xi by auto + moreover have "y \ \s" + using set_rev_mp[OF as(1) s(2)[OF k(1)]] as(2) di i unfolding s mem_interval y_def + by (auto simp: field_simps elim!: ballE[of _ _ i]) ultimately show "y \ \(s - {k})" by auto qed qed hence "\(s - {k}) = {a..b}" unfolding s(6)[THEN sym] by auto hence "{ka \ s - {k}. content ka \ 0} division_of {a..b}" apply-apply(rule assm(2)[rule_format,OF card refl]) @@ -3730,23 +3656,41 @@ "content((\x::'a::ordered_euclidean_space. m *\<^sub>R x + c) ` {a..b}) = (abs m) ^ DIM('a) * content {a..b}" (is "?l = ?r") proof- { presume *:"{a..b}\{} \ ?thesis" show ?thesis apply(cases,rule *,assumption) unfolding not_not using content_empty by auto } - have *:"DIM('a) = card {..{}" show ?thesis proof(cases "m \ 0") - case True show ?thesis unfolding image_affinity_interval if_not_P[OF as] if_P[OF True] - unfolding content_closed_interval'[OF as] apply(subst content_closed_interval') defer apply(subst(2) *) - apply(subst setprod_constant[THEN sym]) apply(rule finite_lessThan) unfolding setprod_timesf[THEN sym] - apply(rule setprod_cong2) using True as unfolding interval_ne_empty euclidean_simps not_le - by(auto simp add:field_simps intro:mult_left_mono) - next case False show ?thesis unfolding image_affinity_interval if_not_P[OF as] if_not_P[OF False] - unfolding content_closed_interval'[OF as] apply(subst content_closed_interval') defer apply(subst(2) *) - apply(subst setprod_constant[THEN sym]) apply(rule finite_lessThan) unfolding setprod_timesf[THEN sym] - apply(rule setprod_cong2) using False as unfolding interval_ne_empty euclidean_simps not_le - by(auto simp add:field_simps mult_le_cancel_left_neg) qed qed + assume as: "{a..b}\{}" + show ?thesis + proof (cases "m \ 0") + case True + with as have "{m *\<^sub>R a + c..m *\<^sub>R b + c} \ {}" + unfolding interval_ne_empty + apply (intro ballI) + apply (erule_tac x=i in ballE) + apply (auto simp: inner_simps intro!: mult_left_mono) + done + moreover from True have *: "\i. (m *\<^sub>R b + c) \ i - (m *\<^sub>R a + c) \ i = m *\<^sub>R (b - a) \ i" + by (simp add: inner_simps field_simps) + ultimately show ?thesis + by (simp add: image_affinity_interval True content_closed_interval' + setprod_timesf setprod_constant inner_diff_left) + next + case False + moreover with as have "{m *\<^sub>R b + c..m *\<^sub>R a + c} \ {}" + unfolding interval_ne_empty + apply (intro ballI) + apply (erule_tac x=i in ballE) + apply (auto simp: inner_simps intro!: mult_left_mono) + done + moreover from False have *: "\i. (m *\<^sub>R a + c) \ i - (m *\<^sub>R b + c) \ i = (-m) *\<^sub>R (b - a) \ i" + by (simp add: inner_simps field_simps) + ultimately show ?thesis + by (simp add: image_affinity_interval content_closed_interval' + setprod_timesf[symmetric] setprod_constant[symmetric] inner_diff_left) + qed +qed lemma has_integral_affinity: fixes a::"'a::ordered_euclidean_space" assumes "(f has_integral i) {a..b}" "m \ 0" shows "((\x. f(m *\<^sub>R x + c)) has_integral ((1 / (abs(m) ^ DIM('a))) *\<^sub>R i)) ((\x. (1 / m) *\<^sub>R x + -((1 / m) *\<^sub>R c)) ` {a..b})" - apply(rule has_integral_twiddle,safe) apply(rule zero_less_power) unfolding euclidean_eq[where 'a='a] - unfolding scaleR_right_distrib euclidean_simps scaleR_scaleR + apply(rule has_integral_twiddle,safe) apply(rule zero_less_power) unfolding euclidean_eq_iff[where 'a='a] + unfolding scaleR_right_distrib inner_simps scaleR_scaleR defer apply(insert assms(2), simp add:field_simps) apply(insert assms(2), simp add:field_simps) apply(rule continuous_intros)+ apply(rule interval_image_affinity_interval)+ apply(rule content_image_affinity_interval) using assms by auto @@ -3757,60 +3701,68 @@ subsection {* Special case of stretching coordinate axes separately. *} lemma image_stretch_interval: - "(\x. \\ k. m k * x$$k) ` {a..b::'a::ordered_euclidean_space} = - (if {a..b} = {} then {} else {(\\ k. min (m(k) * a$$k) (m(k) * b$$k))::'a .. (\\ k. max (m(k) * a$$k) (m(k) * b$$k))})" - (is "?l = ?r") -proof(cases "{a..b}={}") case True thus ?thesis unfolding True by auto -next have *:"\P Q. (\i (\i (\i Q i)" by auto - case False note ab = this[unfolded interval_ne_empty] - show ?thesis apply-apply(rule set_eqI) - proof- fix x::"'a" have **:"\P Q. (\i (\ii ?l \ x \ ?r" unfolding if_not_P[OF False] - unfolding image_iff mem_interval Bex_def euclidean_simps euclidean_eq[where 'a='a] * - unfolding imp_conjR[THEN sym] apply(subst euclidean_lambda_beta'') apply(subst lambda_skolem'[THEN sym]) - apply(rule **,rule,rule) unfolding euclidean_lambda_beta' - proof- fix i assume i:"ixa. (a $$ i \ xa \ xa \ b $$ i) \ x $$ i = m i * xa) = - (min (m i * a $$ i) (m i * b $$ i) \ x $$ i \ x $$ i \ max (m i * a $$ i) (m i * b $$ i))" - proof(cases "m i = 0") case True thus ?thesis using ab i by auto - next case False hence "0 < m i \ 0 > m i" by auto thus ?thesis apply- - proof(erule disjE) assume as:"0 < m i" hence *:"min (m i * a $$ i) (m i * b $$ i) = m i * a $$ i" - "max (m i * a $$ i) (m i * b $$ i) = m i * b $$ i" using ab i unfolding min_def max_def by auto - show ?thesis unfolding * apply rule defer apply(rule_tac x="1 / m i * x$$i" in exI) - using as by(auto simp add:field_simps) - next assume as:"0 > m i" hence *:"max (m i * a $$ i) (m i * b $$ i) = m i * a $$ i" - "min (m i * a $$ i) (m i * b $$ i) = m i * b $$ i" using ab as i unfolding min_def max_def - by(auto simp add:field_simps mult_le_cancel_left_neg intro: order_antisym) - show ?thesis unfolding * apply rule defer apply(rule_tac x="1 / m i * x$$i" in exI) - using as by(auto simp add:field_simps) qed qed qed qed qed - -lemma interval_image_stretch_interval: "\u v. (\x. \\ k. m k * x$$k) ` {a..b::'a::ordered_euclidean_space} = {u..v::'a}" + "(\x. \k\Basis. (m k * (x\k)) *\<^sub>R k) ` {a..b::'a::ordered_euclidean_space} = + (if {a..b} = {} then {} else + {(\k\Basis. (min (m k * (a\k)) (m k * (b\k))) *\<^sub>R k)::'a .. + (\k\Basis. (max (m k * (a\k)) (m k * (b\k))) *\<^sub>R k)})" +proof cases + assume *: "{a..b} \ {}" + show ?thesis + unfolding interval_ne_empty if_not_P[OF *] + apply (simp add: interval image_Collect set_eq_iff euclidean_eq_iff[where 'a='a] ball_conj_distrib[symmetric]) + apply (subst choice_Basis_iff[symmetric]) + proof (intro allI ball_cong refl) + fix x i :: 'a assume "i \ Basis" + with * have a_le_b: "a \ i \ b \ i" + unfolding interval_ne_empty by auto + show "(\xa. x \ i = m i * xa \ a \ i \ xa \ xa \ b \ i) \ + min (m i * (a \ i)) (m i * (b \ i)) \ x \ i \ x \ i \ max (m i * (a \ i)) (m i * (b \ i))" + proof cases + assume "m i \ 0" + moreover then have *: "\a b. a = m i * b \ b = a / m i" + by (auto simp add: field_simps) + moreover have + "min (m i * (a \ i)) (m i * (b \ i)) = (if 0 < m i then m i * (a \ i) else m i * (b \ i))" + "max (m i * (a \ i)) (m i * (b \ i)) = (if 0 < m i then m i * (b \ i) else m i * (a \ i))" + using a_le_b by (auto simp: min_def max_def mult_le_cancel_left) + ultimately show ?thesis using a_le_b + unfolding * by (auto simp add: le_divide_eq divide_le_eq ac_simps) + qed (insert a_le_b, auto) + qed +qed simp + +lemma interval_image_stretch_interval: + "\u v. (\x. \k\Basis. (m k * (x\k))*\<^sub>R k) ` {a..b::'a::ordered_euclidean_space} = {u..v::'a}" unfolding image_stretch_interval by auto lemma content_image_stretch_interval: - "content((\x::'a::ordered_euclidean_space. (\\ k. m k * x$$k)::'a) ` {a..b}) = abs(setprod m {..x::'a::ordered_euclidean_space. (\k\Basis. (m k * (x\k))*\<^sub>R k)::'a) ` {a..b}) = abs(setprod m Basis) * content({a..b})" proof(cases "{a..b} = {}") case True thus ?thesis unfolding content_def image_is_empty image_stretch_interval if_P[OF True] by auto -next case False hence "(\x. (\\ k. m k * x $$ k)::'a) ` {a..b} \ {}" by auto +next case False hence "(\x. (\k\Basis. (m k * (x\k))*\<^sub>R k)::'a) ` {a..b} \ {}" by auto thus ?thesis using False unfolding content_def image_stretch_interval apply- unfolding interval_bounds' if_not_P - unfolding abs_setprod setprod_timesf[THEN sym] apply(rule setprod_cong2) unfolding lessThan_iff euclidean_lambda_beta' - proof- fix i assume i:"i m i > 0) \ m i = 0" by auto - thus "max (m i * a $$ i) (m i * b $$ i) - min (m i * a $$ i) (m i * b $$ i) = \m i\ * (b $$ i - a $$ i)" + unfolding abs_setprod setprod_timesf[THEN sym] apply(rule setprod_cong2) unfolding lessThan_iff + proof (simp only: inner_setsum_left_Basis) + fix i :: 'a assume i:"i\Basis" have "(m i < 0 \ m i > 0) \ m i = 0" by auto + thus "max (m i * (a \ i)) (m i * (b \ i)) - min (m i * (a \ i)) (m i * (b \ i)) = + \m i\ * (b \ i - a \ i)" apply-apply(erule disjE)+ unfolding min_def max_def using False[unfolded interval_ne_empty,rule_format,of i] i by(auto simp add:field_simps not_le mult_le_cancel_left_neg mult_le_cancel_left_pos) qed qed lemma has_integral_stretch: fixes f::"'a::ordered_euclidean_space => 'b::real_normed_vector" - assumes "(f has_integral i) {a..b}" "\kx. f(\\ k. m k * x$$k)) has_integral - ((1/(abs(setprod m {..R i)) ((\x. (\\ k. 1/(m k) * x$$k)::'a) ` {a..b})" + assumes "(f has_integral i) {a..b}" "\k\Basis. ~(m k = 0)" + shows "((\x. f (\k\Basis. (m k * (x\k))*\<^sub>R k)) has_integral + ((1/(abs(setprod m Basis))) *\<^sub>R i)) ((\x. (\k\Basis. (1 / m k * (x\k))*\<^sub>R k)) ` {a..b})" apply(rule has_integral_twiddle[where f=f]) unfolding zero_less_abs_iff content_image_stretch_interval - unfolding image_stretch_interval empty_as_interval euclidean_eq[where 'a='a] using assms -proof- show "\y::'a. continuous (at y) (\x. (\\ k. m k * x $$ k)::'a)" + unfolding image_stretch_interval empty_as_interval euclidean_eq_iff[where 'a='a] using assms +proof- show "\y::'a. continuous (at y) (\x. (\k\Basis. (m k * (x\k))*\<^sub>R k))" apply(rule,rule linear_continuous_at) unfolding linear_linear - unfolding linear_def euclidean_simps euclidean_eq[where 'a='a] by(auto simp add:field_simps) qed auto + unfolding linear_def inner_simps euclidean_eq_iff[where 'a='a] by(auto simp add:field_simps) +qed auto lemma integrable_stretch: fixes f::"'a::ordered_euclidean_space => 'b::real_normed_vector" - assumes "f integrable_on {a..b}" "\kx::'a. f(\\ k. m k * x$$k)) integrable_on ((\x. \\ k. 1/(m k) * x$$k) ` {a..b})" + assumes "f integrable_on {a..b}" "\k\Basis. ~(m k = 0)" + shows "(\x::'a. f (\k\Basis. (m k * (x\k))*\<^sub>R k)) integrable_on ((\x. \k\Basis. (1 / m k * (x\k))*\<^sub>R k) ` {a..b})" using assms unfolding integrable_on_def apply-apply(erule exE) apply(drule has_integral_stretch,assumption) by auto @@ -3856,7 +3808,8 @@ show ?thesis proof(cases,rule *,assumption) assume "\ a < b" hence "a = b" using assms(1) by auto hence *:"{a .. b} = {b}" "f b - f a = 0" by(auto simp add: order_antisym) - show ?thesis unfolding *(2) apply(rule has_integral_null) unfolding content_eq_0 using * `a=b` by auto + show ?thesis unfolding *(2) apply(rule has_integral_null) unfolding content_eq_0 using * `a=b` + by (auto simp: ex_in_conv) qed } assume ab:"a < b" let ?P = "\e. \d. gauge d \ (\p. p tagged_division_of {a..b} \ d fine p \ norm ((\(x, k)\p. content k *\<^sub>R f' x) - (f b - f a)) \ e * content {a..b})" @@ -4106,11 +4059,14 @@ from this(2)[unfolded open_contains_ball,rule_format,OF this(1)] guess k .. note k=conjunctD2[OF this] let ?d = "min k (c - a)/2" show ?thesis apply(rule that[of ?d]) - proof safe show "?d > 0" using k(1) using assms(2) by auto - fix t assume as:"c - ?d < t" "t \ c" let ?thesis = "norm (integral {a..c} f - integral {a..t} f) < e" + proof safe + show "?d > 0" using k(1) using assms(2) by auto + fix t assume as:"c - ?d < t" "t \ c" + let ?thesis = "norm (integral {a..c} f - integral {a..t} f) < e" { presume *:"t < c \ ?thesis" show ?thesis apply(cases "t = c") defer apply(rule *) - apply(subst less_le) using `e>0` as(2) by auto } assume "t < c" + apply(subst less_le) using `e>0` as(2) by auto } + assume "t < c" have "f integrable_on {a..t}" apply(rule integrable_subinterval[OF assms(1)]) using assms(2-3) as(2) by auto from integrable_integral[OF this,unfolded has_integral,rule_format,OF *] guess d2 .. @@ -4123,10 +4079,10 @@ with p(2) have "d2 fine p" unfolding fine_def d3_def apply safe apply(erule_tac x="(a,b)" in ballE)+ by auto note d2_fin = d2(2)[OF conjI[OF p(1) this]] - have *:"{a..c} \ {x. x $$0 \ t} = {a..t}" "{a..c} \ {x. x$$0 \ t} = {t..c}" + have *:"{a..c} \ {x. x \ 1 \ t} = {a..t}" "{a..c} \ {x. x \ 1 \ t} = {t..c}" using assms(2-3) as by(auto simp add:field_simps) have "p \ {(c, {t..c})} tagged_division_of {a..c} \ d1 fine p \ {(c, {t..c})}" apply rule - apply(rule tagged_division_union_interval[of _ _ _ 0 "t"]) unfolding * apply(rule p) + apply(rule tagged_division_union_interval[of _ _ _ 1 "t"]) unfolding * apply(rule p) apply(rule tagged_division_of_self) unfolding fine_def proof safe fix x k y assume "(x,k)\p" "y\k" thus "y\d1 x" using p(2) pt unfolding fine_def d3_def apply- apply(erule_tac x="(x,k)" in ballE)+ by auto @@ -4178,7 +4134,7 @@ { presume *:"a ?thesis" show ?thesis apply(cases,rule *,assumption) proof- case goal1 hence "{a..b} = {x}" using as(1) apply-apply(rule set_eqI) - unfolding atLeastAtMost_iff by(auto simp only:field_simps not_less DIM_real) + unfolding atLeastAtMost_iff by(auto simp only:field_simps not_less) thus ?case using `e>0` by auto qed } assume "a x=b) \ (a x0`, auto) next assume as:"\e>0. ?r e" from this[rule_format,OF zero_less_one] guess C .. note C=conjunctD2[OF this,rule_format] - def c \ "(\\ i. - max B C)::'n::ordered_euclidean_space" and d \ "(\\ i. max B C)::'n::ordered_euclidean_space" + def c \ "(\i\Basis. (- max B C) *\<^sub>R i)::'n::ordered_euclidean_space" + def d \ "(\i\Basis. max B C *\<^sub>R i)::'n::ordered_euclidean_space" have c_d:"{a..b} \ {c..d}" apply safe apply(drule B(2)) unfolding mem_interval - proof case goal1 thus ?case using component_le_norm[of x i] unfolding c_def d_def - by(auto simp add:field_simps) qed + proof + case goal1 thus ?case using Basis_le_norm[OF `i\Basis`, of x] unfolding c_def d_def + by(auto simp add:field_simps setsum_negf) + qed have "ball 0 C \ {c..d}" apply safe unfolding mem_interval mem_ball dist_norm - proof case goal1 thus ?case using component_le_norm[of x i] unfolding c_def d_def by auto qed + proof + case goal1 thus ?case + using Basis_le_norm[OF `i\Basis`, of x] unfolding c_def d_def by (auto simp: setsum_negf) + qed from C(2)[OF this] have "\y. (f has_integral y) {a..b}" unfolding has_integral_restrict_closed_subintervals_eq[OF c_d,THEN sym] unfolding s by auto then guess y .. note y=this have "y = i" proof(rule ccontr) assume "y\i" hence "0 < norm (y - i)" by auto from as[rule_format,OF this] guess C .. note C=conjunctD2[OF this,rule_format] - def c \ "(\\ i. - max B C)::'n::ordered_euclidean_space" and d \ "(\\ i. max B C)::'n::ordered_euclidean_space" + def c \ "(\i\Basis. (- max B C) *\<^sub>R i)::'n::ordered_euclidean_space" + def d \ "(\i\Basis. max B C *\<^sub>R i)::'n::ordered_euclidean_space" have c_d:"{a..b} \ {c..d}" apply safe apply(drule B(2)) unfolding mem_interval - proof case goal1 thus ?case using component_le_norm[of x i] unfolding c_def d_def - by(auto simp add:field_simps) qed + proof case goal1 thus ?case using Basis_le_norm[of i x] unfolding c_def d_def + by(auto simp add:field_simps setsum_negf) qed have "ball 0 C \ {c..d}" apply safe unfolding mem_interval mem_ball dist_norm - proof case goal1 thus ?case using component_le_norm[of x i] unfolding c_def d_def by auto qed + proof case goal1 thus ?case using Basis_le_norm[of i x] unfolding c_def d_def by (auto simp: setsum_negf) qed note C(2)[OF this] then guess z .. note z = conjunctD2[OF this, unfolded s] note this[unfolded has_integral_restrict_closed_subintervals_eq[OF c_d]] hence "z = y" "norm (z - i) < norm (y - i)" apply- apply(rule has_integral_unique[OF _ y(1)]) . thus False by auto qed thus ?l using y unfolding s by auto qed qed -(*lemma has_integral_trans[simp]: fixes f::"'n::ordered_euclidean_space \ real" shows - "((\x. vec1 (f x)) has_integral vec1 i) s \ (f has_integral i) s" - unfolding has_integral'[unfolded has_integral] -proof case goal1 thus ?case apply safe - apply(erule_tac x=e in allE,safe) apply(rule_tac x=B in exI,safe) - apply(erule_tac x=a in allE, erule_tac x=b in allE,safe) - apply(rule_tac x="dest_vec1 z" in exI,safe) apply(erule_tac x=ea in allE,safe) - apply(rule_tac x=d in exI,safe) apply(erule_tac x=p in allE,safe) - apply(subst(asm)(2) norm_vector_1) unfolding split_def - unfolding setsum_component Cart_nth.diff cond_value_iff[of dest_vec1] - Cart_nth.scaleR vec1_dest_vec1 zero_index apply assumption - apply(subst(asm)(2) norm_vector_1) by auto -next case goal2 thus ?case apply safe - apply(erule_tac x=e in allE,safe) apply(rule_tac x=B in exI,safe) - apply(erule_tac x=a in allE, erule_tac x=b in allE,safe) - apply(rule_tac x="vec1 z" in exI,safe) apply(erule_tac x=ea in allE,safe) - apply(rule_tac x=d in exI,safe) apply(erule_tac x=p in allE,safe) - apply(subst norm_vector_1) unfolding split_def - unfolding setsum_component Cart_nth.diff cond_value_iff[of dest_vec1] - Cart_nth.scaleR vec1_dest_vec1 zero_index apply assumption - apply(subst norm_vector_1) by auto qed - -lemma integral_trans[simp]: assumes "(f::'n::ordered_euclidean_space \ real) integrable_on s" - shows "integral s (\x. vec1 (f x)) = vec1 (integral s f)" - apply(rule integral_unique) using assms by auto - -lemma integrable_on_trans[simp]: fixes f::"'n::ordered_euclidean_space \ real" shows - "(\x. vec1 (f x)) integrable_on s \ (f integrable_on s)" - unfolding integrable_on_def - apply(subst(2) vec1_dest_vec1(1)[THEN sym]) unfolding has_integral_trans - apply safe defer apply(rule_tac x="vec1 y" in exI) by auto *) - lemma has_integral_le: fixes f::"'n::ordered_euclidean_space \ real" assumes "(f has_integral i) s" "(g has_integral j) s" "\x\s. (f x) \ (g x)" - shows "i \ j" using has_integral_component_le[OF assms(1-2), of 0] using assms(3) by auto + shows "i \ j" + using has_integral_component_le[OF _ assms(1-2), of 1] using assms(3) by auto lemma integral_le: fixes f::"'n::ordered_euclidean_space \ real" assumes "f integrable_on s" "g integrable_on s" "\x\s. f x \ g x" @@ -4422,7 +4354,7 @@ lemma has_integral_nonneg: fixes f::"'n::ordered_euclidean_space \ real" assumes "(f has_integral i) s" "\x\s. 0 \ f x" shows "0 \ i" - using has_integral_component_nonneg[of "f" "i" s 0] + using has_integral_component_nonneg[of 1 f i s] unfolding o_def using assms by auto lemma integral_nonneg: fixes f::"'n::ordered_euclidean_space \ real" @@ -4519,19 +4451,20 @@ subsection {* More lemmas that are useful later. *} lemma has_integral_subset_component_le: fixes f::"'n::ordered_euclidean_space \ 'm::ordered_euclidean_space" - assumes "s \ t" "(f has_integral i) s" "(f has_integral j) t" "\x\t. 0 \ f(x)$$k" - shows "i$$k \ j$$k" + assumes k: "k\Basis" and as: "s \ t" "(f has_integral i) s" "(f has_integral j) t" "\x\t. 0 \ f(x)\k" + shows "i\k \ j\k" proof- note has_integral_restrict_univ[THEN sym, of f] - note assms(2-3)[unfolded this] note * = has_integral_component_le[OF this] - show ?thesis apply(rule *) using assms(1,4) by auto qed + note as(2-3)[unfolded this] note * = has_integral_component_le[OF k this] + show ?thesis apply(rule *) using as(1,4) by auto qed lemma has_integral_subset_le: fixes f::"'n::ordered_euclidean_space \ real" - assumes "s \ t" "(f has_integral i) s" "(f has_integral j) t" "\x\t. 0 \ f(x)" - shows "i \ j" using has_integral_subset_component_le[OF assms(1), of "f" "i" "j" 0] using assms by auto + assumes as: "s \ t" "(f has_integral i) s" "(f has_integral j) t" "\x\t. 0 \ f(x)" + shows "i \ j" + using has_integral_subset_component_le[OF _ assms(1), of 1 f i j] using assms by auto lemma integral_subset_component_le: fixes f::"'n::ordered_euclidean_space \ 'm::ordered_euclidean_space" - assumes "s \ t" "f integrable_on s" "f integrable_on t" "\x \ t. 0 \ f(x)$$k" - shows "(integral s f)$$k \ (integral t f)$$k" + assumes "k\Basis" "s \ t" "f integrable_on s" "f integrable_on t" "\x \ t. 0 \ f(x)\k" + shows "(integral s f)\k \ (integral t f)\k" apply(rule has_integral_subset_component_le) using assms by auto lemma integral_subset_le: fixes f::"'n::ordered_euclidean_space \ real" @@ -4553,13 +4486,13 @@ let ?f = "\x. if x \ s then f x else 0" show ?r proof safe fix a b::"'n::ordered_euclidean_space" from as[OF zero_less_one] guess B .. note B=conjunctD2[OF this,rule_format] - let ?a = "(\\ i. min (a$$i) (-B))::'n::ordered_euclidean_space" and ?b = "(\\ i. max (b$$i) B)::'n::ordered_euclidean_space" + let ?a = "\i\Basis. min (a\i) (-B) *\<^sub>R i::'n" and ?b = "\i\Basis. max (b\i) B *\<^sub>R i::'n" show "?f integrable_on {a..b}" apply(rule integrable_subinterval[of _ ?a ?b]) proof- have "ball 0 B \ {?a..?b}" apply safe unfolding mem_ball mem_interval dist_norm - proof case goal1 thus ?case using component_le_norm[of x i] by(auto simp add:field_simps) qed + proof case goal1 thus ?case using Basis_le_norm[of i x] by(auto simp add:field_simps) qed from B(2)[OF this] guess z .. note conjunct1[OF this] thus "?f integrable_on {?a..?b}" unfolding integrable_on_def by auto - show "{a..b} \ {?a..?b}" apply safe unfolding mem_interval apply(rule,erule_tac x=i in allE) by auto qed + show "{a..b} \ {?a..?b}" apply safe unfolding mem_interval apply(rule,erule_tac x=i in ballE) by auto qed fix e::real assume "e>0" from as[OF this] guess B .. note B=conjunctD2[OF this,rule_format] show "\B>0. \a b. ball 0 B \ {a..b} \ @@ -4584,14 +4517,15 @@ using B(2)[OF goal1(1)] B(2)[OF goal1(2)] by auto qed qed next assume ?r note as = conjunctD2[OF this,rule_format] - have "Cauchy (\n. integral ({(\\ i. - real n)::'n .. (\\ i. real n)}) (\x. if x \ s then f x else 0))" + let ?cube = "\n. {(\i\Basis. - real n *\<^sub>R i)::'n .. (\i\Basis. real n *\<^sub>R i)} :: 'n set" + have "Cauchy (\n. integral (?cube n) (\x. if x \ s then f x else 0))" proof(unfold Cauchy_def,safe) case goal1 from as(2)[OF this] guess B .. note B = conjunctD2[OF this,rule_format] from real_arch_simple[of B] guess N .. note N = this - { fix n assume n:"n \ N" have "ball 0 B \ {(\\ i. - real n)::'n..\\ i. real n}" apply safe + { fix n assume n:"n \ N" have "ball 0 B \ ?cube n" apply safe unfolding mem_ball mem_interval dist_norm - proof case goal1 thus ?case using component_le_norm[of x i] - using n N by(auto simp add:field_simps) qed } + proof case goal1 thus ?case using Basis_le_norm[of i x] `i\Basis` + using n N by(auto simp add:field_simps setsum_negf) qed } thus ?case apply-apply(rule_tac x=N in exI) apply safe unfolding dist_norm apply(rule B(2)) by auto qed from this[unfolded convergent_eq_cauchy[THEN sym]] guess i .. note i = this[THEN LIMSEQ_D] @@ -4611,9 +4545,9 @@ proof safe show "N \ n" using n by auto fix x::"'n::ordered_euclidean_space" assume x:"x \ ball 0 B" hence "x\ ball 0 ?B" by auto thus "x\{a..b}" using ab by blast - show "x\{\\ i. - real n..\\ i. real n}" using x unfolding mem_interval mem_ball dist_norm apply- - proof case goal1 thus ?case using component_le_norm[of x i] - using n by(auto simp add:field_simps) qed qed qed qed qed + show "x\?cube n" using x unfolding mem_interval mem_ball dist_norm apply- + proof case goal1 thus ?case using Basis_le_norm[of i x] `i\Basis` + using n by(auto simp add:field_simps setsum_negf) qed qed qed qed qed lemma integrable_altD: fixes f::"'n::ordered_euclidean_space \ 'a::banach" assumes "f integrable_on s" @@ -4676,11 +4610,13 @@ note g = this(1) and this(2)[OF *] from this(2) guess B1 .. note B1 = conjunctD2[OF this,rule_format] note obt(2)[unfolded has_integral_alt'[of h]] note conjunctD2[OF this, rule_format] note h = this(1) and this(2)[OF *] from this(2) guess B2 .. note B2 = conjunctD2[OF this,rule_format] - def c \ "(\\ i. min (a$$i) (- (max B1 B2)))::'n" and d \ "(\\ i. max (b$$i) (max B1 B2))::'n" - have *:"ball 0 B1 \ {c..d}" "ball 0 B2 \ {c..d}" apply safe unfolding mem_ball mem_interval dist_norm - proof(rule_tac[!] allI) - case goal1 thus ?case using component_le_norm[of x i] unfolding c_def d_def by auto next - case goal2 thus ?case using component_le_norm[of x i] unfolding c_def d_def by auto qed + def c \ "\i\Basis. min (a\i) (- (max B1 B2)) *\<^sub>R i::'n" + def d \ "\i\Basis. max (b\i) (max B1 B2) *\<^sub>R i::'n" + have *:"ball 0 B1 \ {c..d}" "ball 0 B2 \ {c..d}" + apply safe unfolding mem_ball mem_interval dist_norm + proof(rule_tac[!] ballI) + case goal1 thus ?case using Basis_le_norm[of i x] unfolding c_def d_def by auto next + case goal2 thus ?case using Basis_le_norm[of i x] unfolding c_def d_def by auto qed have **:"\ch cg ag ah::real. norm(ah - ag) \ norm(ch - cg) \ norm(cg - i) < e / 4 \ norm(ch - j) < e / 4 \ norm(ag - ah) < e" using obt(3) unfolding real_norm_def by arith @@ -4700,7 +4636,7 @@ unfolding integral_sub[OF h g,THEN sym] real_norm_def apply(subst **) defer apply(subst **) defer apply(rule has_integral_subset_le) defer apply(rule integrable_integral integrable_sub h g)+ proof safe fix x assume "x\{a..b}" thus "x\{c..d}" unfolding mem_interval c_def d_def - apply - apply rule apply(erule_tac x=i in allE) by auto + apply - apply rule apply(erule_tac x=i in ballE) by auto qed(insert obt(4), auto) qed(insert obt(4), auto) qed note interv = this show ?thesis unfolding integrable_alt[of f] apply safe apply(rule interv) @@ -5019,7 +4955,7 @@ proof(case_tac[!] "content {a..b} = 0") assume as:"content {a..b} = 0" show ?thesis using integrable_on_null[OF as] unfolding integral_null[OF as] using tendsto_const by auto next assume ab:"content {a..b} \ 0" - have fg:"\x\{a..b}. \ k. (f k x) $$ 0 \ (g x) $$ 0" + have fg:"\x\{a..b}. \ k. (f k x) \ 1 \ (g x) \ 1" proof safe case goal1 note assms(3)[rule_format,OF this] note * = Lim_component_ge[OF this trivial_limit_sequentially] show ?case apply(rule *) unfolding eventually_sequentially @@ -5030,10 +4966,10 @@ apply rule apply(rule integral_le) apply safe apply(rule assms(1-2)[rule_format])+ using assms(4) by auto then guess i .. note i=this - have i':"\k. (integral({a..b}) (f k)) \ i$$0" + have i':"\k. (integral({a..b}) (f k)) \ i\1" apply(rule Lim_component_ge,rule i) apply(rule trivial_limit_sequentially) unfolding eventually_sequentially apply(rule_tac x=k in exI) - apply(rule transitive_stepwise_le) prefer 3 unfolding Eucl_real_simps apply(rule integral_le) + apply(rule transitive_stepwise_le) prefer 3 unfolding inner_simps real_inner_1_right apply(rule integral_le) apply(rule assms(1-2)[rule_format])+ using assms(2) by auto have "(g has_integral i) {a..b}" unfolding has_integral @@ -5044,7 +4980,7 @@ apply(rule divide_pos_pos) by auto from choice[OF this] guess c .. note c=conjunctD2[OF this[rule_format],rule_format] - have "\r. \k\r. 0 \ i$$0 - (integral {a..b} (f k)) \ i$$0 - (integral {a..b} (f k)) < e / 4" + have "\r. \k\r. 0 \ i\1 - (integral {a..b} (f k)) \ i\1 - (integral {a..b} (f k)) < e / 4" proof- case goal1 have "e/4 > 0" using e by auto from LIMSEQ_D [OF i this] guess r .. thus ?case apply(rule_tac x=r in exI) apply rule @@ -5052,14 +4988,15 @@ proof- case goal1 thus ?case using i'[of k] by auto qed qed then guess r .. note r=conjunctD2[OF this[rule_format]] - have "\x\{a..b}. \n\r. \k\n. 0 \ (g x)$$0 - (f k x)$$0 \ - (g x)$$0 - (f k x)$$0 < e / (4 * content({a..b}))" + have "\x\{a..b}. \n\r. \k\n. 0 \ (g x)\1 - (f k x)\1 \ + (g x)\1 - (f k x)\1 < e / (4 * content({a..b}))" proof case goal1 have "e / (4 * content {a..b}) > 0" apply(rule divide_pos_pos,fact) using ab content_pos_le[of a b] by auto from assms(3)[rule_format, OF goal1, THEN LIMSEQ_D, OF this] guess n .. note n=this thus ?case apply(rule_tac x="n + r" in exI) apply safe apply(erule_tac[2-3] x=k in allE) - unfolding dist_real_def using fg[rule_format,OF goal1] by(auto simp add:field_simps) qed + unfolding dist_real_def using fg[rule_format,OF goal1] + by (auto simp add:field_simps) qed from bchoice[OF this] guess m .. note m=conjunctD2[OF this[rule_format],rule_format] def d \ "\x. c (m x) x" @@ -5118,14 +5055,14 @@ next case goal3 note comb = integral_combine_tagged_division_topdown[OF assms(1)[rule_format] p(1)] - have *:"\sr sx ss ks kr::real. kr = sr \ ks = ss \ ks \ i \ sr \ sx \ sx \ ss \ 0 \ i$$0 - kr$$0 - \ i$$0 - kr$$0 < e/4 \ abs(sx - i) < e/4" by auto + have *:"\sr sx ss ks kr::real. kr = sr \ ks = ss \ ks \ i \ sr \ sx \ sx \ ss \ 0 \ i\1 - kr\1 + \ i\1 - kr\1 < e/4 \ abs(sx - i) < e/4" by auto show ?case unfolding real_norm_def apply(rule *[rule_format],safe) - apply(rule comb[of r],rule comb[of s]) apply(rule i'[unfolded Eucl_real_simps]) + apply(rule comb[of r],rule comb[of s]) apply(rule i'[unfolded real_inner_1_right]) apply(rule_tac[1-2] setsum_mono) unfolding split_paired_all split_conv apply(rule_tac[1-2] integral_le[OF ]) - proof safe show "0 \ i$$0 - (integral {a..b} (f r))$$0" using r(1) by auto - show "i$$0 - (integral {a..b} (f r))$$0 < e / 4" using r(2) by auto + proof safe show "0 \ i\1 - (integral {a..b} (f r))\1" using r(1) by auto + show "i\1 - (integral {a..b} (f r))\1 < e / 4" using r(2) by auto fix x k assume xk:"(x,k)\p" from p'(4)[OF this] guess u v apply-by(erule exE)+ note uv=this show "f r integrable_on k" "f s integrable_on k" "f (m x) integrable_on k" "f (m x) integrable_on k" unfolding uv apply(rule_tac[!] integrable_on_subinterval[OF assms(1)[rule_format]]) @@ -5147,7 +5084,7 @@ \k. \x\s. (f k x) \ (f (Suc k) x) \ \x\s. ((\k. f k x) ---> g x) sequentially \ bounded {integral s (f k)| k. True} \ g integrable_on s \ ((\k. integral s (f k)) ---> integral s g) sequentially" proof- case goal1 note assms=this[rule_format] - have "\x\s. \k. (f k x)$$0 \ (g x)$$0" apply safe apply(rule Lim_component_ge) + have "\x\s. \k. (f k x)\1 \ (g x)\1" apply safe apply(rule Lim_component_ge) apply(rule goal1(4)[rule_format],assumption) apply(rule trivial_limit_sequentially) unfolding eventually_sequentially apply(rule_tac x=k in exI) apply(rule transitive_stepwise_le) using goal1(3) by auto note fg=this[rule_format] @@ -5156,9 +5093,10 @@ apply(rule goal1(5)) apply(rule,rule integral_le) apply(rule goal1(2)[rule_format])+ using goal1(3) by auto then guess i .. note i=this have "\k. \x\s. \n\k. f k x \ f n x" apply(rule,rule transitive_stepwise_le) using goal1(3) by auto - hence i':"\k. (integral s (f k))$$0 \ i$$0" apply-apply(rule,rule Lim_component_ge) + hence i':"\k. (integral s (f k))\1 \ i\1" apply-apply(rule,rule Lim_component_ge) apply(rule i,rule trivial_limit_sequentially) unfolding eventually_sequentially apply(rule_tac x=k in exI,safe) apply(rule integral_component_le) + apply simp apply(rule goal1(2)[rule_format])+ by auto note int = assms(2)[unfolded integrable_alt[of _ s],THEN conjunct1,rule_format] @@ -5199,14 +5137,14 @@ apply(rule norm_triangle_half_l) using B(2)[rule_format,OF ab] N[rule_format,of N] apply-defer apply(subst norm_minus_commute) by auto have *:"\f1 f2 g. abs(f1 - i) < e / 2 \ abs(f2 - g) < e / 2 \ f1 \ f2 \ f2 \ i - \ abs(g - i) < e" unfolding Eucl_real_simps by arith + \ abs(g - i) < e" unfolding real_inner_1_right by arith show "norm (integral {a..b} (\x. if x \ s then g x else 0) - i) < e" unfolding real_norm_def apply(rule *[rule_format]) apply(rule **[unfolded real_norm_def]) apply(rule M[rule_format,of "M + N",unfolded real_norm_def]) apply(rule le_add1) apply(rule integral_le[OF int int]) defer - apply(rule order_trans[OF _ i'[rule_format,of "M + N",unfolded Eucl_real_simps]]) - proof safe case goal2 have "\m. x\s \ \n\m. (f m x)$$0 \ (f n x)$$0" + apply(rule order_trans[OF _ i'[rule_format,of "M + N",unfolded real_inner_1_right]]) + proof safe case goal2 have "\m. x\s \ \n\m. (f m x)\1 \ (f n x)\1" apply(rule transitive_stepwise_le) using assms(3) by auto thus ?case by auto next case goal1 show ?case apply(subst integral_restrict_univ[THEN sym,OF int]) unfolding ifif integral_restrict_univ[OF int'] @@ -5323,9 +5261,9 @@ lemma integral_norm_bound_integral_component: fixes f::"'n::ordered_euclidean_space \ 'a::banach" fixes g::"'n => 'b::ordered_euclidean_space" - assumes "f integrable_on s" "g integrable_on s" "\x\s. norm(f x) \ (g x)$$k" - shows "norm(integral s f) \ (integral s g)$$k" -proof- have "norm (integral s f) \ integral s ((\x. x $$ k) o g)" + assumes "f integrable_on s" "g integrable_on s" "\x\s. norm(f x) \ (g x)\k" + shows "norm(integral s f) \ (integral s g)\k" +proof- have "norm (integral s f) \ integral s ((\x. x \ k) o g)" apply(rule integral_norm_bound_integral[OF assms(1)]) apply(rule integrable_linear[OF assms(2)],rule) unfolding o_def by(rule assms) @@ -5333,8 +5271,8 @@ lemma has_integral_norm_bound_integral_component: fixes f::"'n::ordered_euclidean_space \ 'a::banach" fixes g::"'n => 'b::ordered_euclidean_space" - assumes "(f has_integral i) s" "(g has_integral j) s" "\x\s. norm(f x) \ (g x)$$k" - shows "norm(i) \ j$$k" using integral_norm_bound_integral_component[of f s g k] + assumes "(f has_integral i) s" "(g has_integral j) s" "\x\s. norm(f x) \ (g x)\k" + shows "norm(i) \ j\k" using integral_norm_bound_integral_component[of f s g k] unfolding integral_unique[OF assms(1)] integral_unique[OF assms(2)] using assms by auto @@ -5742,97 +5680,126 @@ shows "(\x. setsum (\a. f a x) t) absolutely_integrable_on s" using assms(1,2) apply induct defer apply(subst setsum.insert) apply assumption+ by(rule,auto) -lemma absolutely_integrable_vector_abs: fixes f::"'a::ordered_euclidean_space => 'b::ordered_euclidean_space" - assumes "f absolutely_integrable_on s" - shows "(\x. (\\ i. abs(f x$$i))::'c::ordered_euclidean_space) absolutely_integrable_on s" -proof- have *:"\x. ((\\ i. abs(f x$$i))::'c::ordered_euclidean_space) = (setsum (\i. - (((\y. (\\ j. if j = i then y else 0)) o - (((\x. (norm((\\ j. if j = i then x$$i else 0)::'c::ordered_euclidean_space))) o f))) x)) {..i xa. ((if i = xa then f x $$ xa else 0) * (if i = xa then f x $$ xa else 0)) = - (if i = xa then (f x $$ xa) * (f x $$ xa) else 0)" by auto - have *:"\xa. norm ((\\ j. if j = xa then f x $$ xa else 0)::'c) = (if xaxay. (\\ j. if j = xa then y else 0)::'c) \ - (\x. (norm ((\\ j. if j = xa then x $$ xa else 0)::'c))) \ f) x $$ i)" - unfolding o_def * apply(rule setsum_cong2) - unfolding euclidean_lambda_beta'[OF goal1 ] by auto - finally show ?case unfolding o_def . qed - show ?thesis unfolding * apply(rule absolutely_integrable_setsum) apply(rule finite_lessThan) - apply(rule absolutely_integrable_linear) unfolding o_def apply(rule absolutely_integrable_norm) - apply(rule absolutely_integrable_linear[OF assms,unfolded o_def]) unfolding linear_linear - apply(rule_tac[!] linearI) unfolding euclidean_eq[where 'a='c] - by(auto simp:euclidean_component_scaleR[where 'a=real,unfolded real_scaleR_def]) +lemma bounded_linear_setsum: + fixes f :: "'i \ 'a::real_normed_vector \ 'b::real_normed_vector" + assumes f: "\i. i\I \ bounded_linear (f i)" + shows "bounded_linear (\x. \i\I. f i x)" +proof cases + assume "finite I" from this f show ?thesis + by (induct I) (auto intro!: bounded_linear_add bounded_linear_zero) +qed (simp add: bounded_linear_zero) + +lemma absolutely_integrable_vector_abs: + fixes f::"'a::ordered_euclidean_space => 'b::ordered_euclidean_space" + fixes T :: "'c::ordered_euclidean_space \ 'b" + assumes f: "f absolutely_integrable_on s" + shows "(\x. (\i\Basis. abs(f x\T i) *\<^sub>R i)) absolutely_integrable_on s" + (is "?Tf absolutely_integrable_on s") +proof - + have if_distrib: "\P A B x. (if P then A else B) *\<^sub>R x = (if P then A *\<^sub>R x else B *\<^sub>R x)" + by simp + have *: "\x. ?Tf x = (\i\Basis. + ((\y. (\j\Basis. (if j = i then y else 0) *\<^sub>R j)) o + (\x. (norm (\j\Basis. (if j = i then f x\T i else 0) *\<^sub>R j)))) x)" + by (simp add: comp_def if_distrib setsum_cases) + show ?thesis + unfolding * + apply (rule absolutely_integrable_setsum[OF finite_Basis]) + apply (rule absolutely_integrable_linear) + apply (rule absolutely_integrable_norm) + apply (rule absolutely_integrable_linear[OF f, unfolded o_def]) + apply (auto simp: linear_linear euclidean_eq_iff[where 'a='c] inner_simps intro!: linearI) + done qed -lemma absolutely_integrable_max: fixes f g::"'m::ordered_euclidean_space \ 'n::ordered_euclidean_space" +lemma absolutely_integrable_max: + fixes f g::"'m::ordered_euclidean_space \ 'n::ordered_euclidean_space" assumes "f absolutely_integrable_on s" "g absolutely_integrable_on s" - shows "(\x. (\\ i. max (f(x)$$i) (g(x)$$i))::'n::ordered_euclidean_space) absolutely_integrable_on s" -proof- have *:"\x. (1 / 2) *\<^sub>R (((\\ i. \(f x - g x) $$ i\)::'n) + (f x + g x)) = (\\ i. max (f(x)$$i) (g(x)$$i))" - unfolding euclidean_eq[where 'a='n] by auto + shows "(\x. (\i\Basis. max (f(x)\i) (g(x)\i) *\<^sub>R i)::'n) absolutely_integrable_on s" +proof - + have *:"\x. (1 / 2) *\<^sub>R (((\i\Basis. \(f x - g x) \ i\ *\<^sub>R i)::'n) + (f x + g x)) = + (\i\Basis. max (f(x)\i) (g(x)\i) *\<^sub>R i)" + unfolding euclidean_eq_iff[where 'a='n] by (auto simp: inner_simps) note absolutely_integrable_sub[OF assms] absolutely_integrable_add[OF assms] - note absolutely_integrable_vector_abs[OF this(1)] this(2) - note absolutely_integrable_add[OF this] note absolutely_integrable_cmul[OF this,of "1/2"] - thus ?thesis unfolding * . qed - -lemma absolutely_integrable_min: fixes f g::"'m::ordered_euclidean_space \ 'n::ordered_euclidean_space" + note absolutely_integrable_vector_abs[OF this(1), where T="\x. x"] this(2) + note absolutely_integrable_add[OF this] + note absolutely_integrable_cmul[OF this, of "1/2"] + thus ?thesis unfolding * . +qed + +lemma absolutely_integrable_min: + fixes f g::"'m::ordered_euclidean_space \ 'n::ordered_euclidean_space" assumes "f absolutely_integrable_on s" "g absolutely_integrable_on s" - shows "(\x. (\\ i. min (f(x)$$i) (g(x)$$i))::'n::ordered_euclidean_space) absolutely_integrable_on s" -proof- have *:"\x. (1 / 2) *\<^sub>R ((f x + g x) - ((\\ i. \(f x - g x) $$ i\)::'n)) = (\\ i. min (f(x)$$i) (g(x)$$i))" - unfolding euclidean_eq[where 'a='n] by auto + shows "(\x. (\i\Basis. min (f(x)\i) (g(x)\i) *\<^sub>R i)::'n) absolutely_integrable_on s" +proof - + have *:"\x. (1 / 2) *\<^sub>R ((f x + g x) - (\i\Basis. \(f x - g x) \ i\ *\<^sub>R i::'n)) = + (\i\Basis. min (f(x)\i) (g(x)\i) *\<^sub>R i)" + unfolding euclidean_eq_iff[where 'a='n] by (auto simp: inner_simps) + note absolutely_integrable_add[OF assms] absolutely_integrable_sub[OF assms] - note this(1) absolutely_integrable_vector_abs[OF this(2)] - note absolutely_integrable_sub[OF this] note absolutely_integrable_cmul[OF this,of "1/2"] - thus ?thesis unfolding * . qed - -lemma absolutely_integrable_abs_eq: fixes f::"'n::ordered_euclidean_space \ 'm::ordered_euclidean_space" + note this(1) absolutely_integrable_vector_abs[OF this(2), where T="\x. x"] + note absolutely_integrable_sub[OF this] + note absolutely_integrable_cmul[OF this,of "1/2"] + thus ?thesis unfolding * . +qed + +lemma absolutely_integrable_abs_eq: + fixes f::"'n::ordered_euclidean_space \ 'm::ordered_euclidean_space" shows "f absolutely_integrable_on s \ f integrable_on s \ - (\x. (\\ i. abs(f x$$i))::'m) integrable_on s" (is "?l = ?r") -proof assume ?l thus ?r apply-apply rule defer + (\x. (\i\Basis. abs(f x\i) *\<^sub>R i)::'m) integrable_on s" (is "?l = ?r") +proof + assume ?l thus ?r apply-apply rule defer apply(drule absolutely_integrable_vector_abs) by auto -next assume ?r { presume lem:"\f::'n \ 'm. f integrable_on UNIV \ - (\x. (\\ i. abs(f(x)$$i))::'m) integrable_on UNIV \ f absolutely_integrable_on UNIV" - have *:"\x. (\\ i. \(if x \ s then f x else 0) $$ i\) = (if x\s then (\\ i. \f x $$ i\) else (0::'m))" - unfolding euclidean_eq[where 'a='m] by auto +next + assume ?r + { presume lem:"\f::'n \ 'm. f integrable_on UNIV \ + (\x. (\i\Basis. \f x\i\ *\<^sub>R i)::'m) integrable_on UNIV \ f absolutely_integrable_on UNIV" + have *:"\x. (\i\Basis. \(if x \ s then f x else 0) \ i\ *\<^sub>R i) = + (if x\s then (\i\Basis. \f x \ i\ *\<^sub>R i) else (0::'m))" + unfolding euclidean_eq_iff[where 'a='m] by auto show ?l apply(subst absolutely_integrable_restrict_univ[THEN sym]) apply(rule lem) unfolding integrable_restrict_univ * using `?r` by auto } - fix f::"'n::ordered_euclidean_space \ 'm::ordered_euclidean_space" assume assms:"f integrable_on UNIV" - "(\x. (\\ i. abs(f(x)$$i))::'m::ordered_euclidean_space) integrable_on UNIV" - let ?B = "setsum (\i. integral UNIV (\x. (\\ j. abs(f x$$j)) ::'m::ordered_euclidean_space) $$ i) {.. 'm::ordered_euclidean_space" + assume assms:"f integrable_on UNIV" "(\x. (\i\Basis. \f x\i\ *\<^sub>R i)::'m) integrable_on UNIV" + let ?B = "\i\Basis. integral UNIV (\x. (\i\Basis. \f x\i\ *\<^sub>R i)::'m) \ i" show "f absolutely_integrable_on UNIV" apply(rule bounded_variation_absolutely_integrable[OF assms(1), where B="?B"],safe) proof- case goal1 note d=this and d'=division_ofD[OF this] have "(\k\d. norm (integral k f)) \ - (\k\d. setsum (op $$ (integral k (\x. (\\ j. \f x $$ j\)::'m))) {..k\d. setsum (op \ (integral k (\x. (\i\Basis. \f x\i\ *\<^sub>R i)::'m))) Basis)" + apply(rule setsum_mono) apply(rule order_trans[OF norm_le_l1]) apply(rule setsum_mono) unfolding lessThan_iff - proof- fix k and i assume "k\d" and i:"id" and i:"i\Basis" from d'(4)[OF this(1)] guess a b apply-by(erule exE)+ note ab=this - show "\integral k f $$ i\ \ integral k (\x. (\\ j. \f x $$ j\)::'m) $$ i" apply(rule abs_leI) - unfolding euclidean_component_minus[THEN sym] defer apply(subst integral_neg[THEN sym]) - defer apply(rule_tac[1-2] integral_component_le) apply(rule integrable_neg) + show "\integral k f \ i\ \ integral k (\x. (\i\Basis. \f x\i\ *\<^sub>R i)::'m) \ i" + apply (rule abs_leI) + unfolding inner_minus_left[THEN sym] defer apply(subst integral_neg[THEN sym]) + defer apply(rule_tac[1-2] integral_component_le[OF i]) apply(rule integrable_neg) using integrable_on_subinterval[OF assms(1),of a b] - integrable_on_subinterval[OF assms(2),of a b] unfolding ab by auto - qed also have "... \ setsum (op $$ (integral UNIV (\x. (\\ j. \f x $$ j\))::'m)) {.. setsum (op \ (integral UNIV (\x. (\i\Basis. \f x\i\ *\<^sub>R i)::'m))) Basis" apply(subst setsum_commute,rule setsum_mono) - proof- case goal1 have *:"(\x. (\\ j. \f x $$ j\)::'m) integrable_on \d" + proof- case goal1 have *:"(\x. \i\Basis. \f x\i\ *\<^sub>R i::'m) integrable_on \d" using integrable_on_subdivision[OF d assms(2)] by auto - have "(\i\d. integral i (\x. (\\ j. \f x $$ j\)::'m) $$ j) - = integral (\d) (\x. (\\ j. abs(f x$$j)) ::'m::ordered_euclidean_space) $$ j" - unfolding euclidean_component_setsum[THEN sym] integral_combine_division_topdown[OF * d] .. - also have "... \ integral UNIV (\x. (\\ j. \f x $$ j\)::'m) $$ j" - apply(rule integral_subset_component_le) using assms * by auto + have "(\i\d. integral i (\x. \i\Basis. \f x\i\ *\<^sub>R i::'m) \ j) + = integral (\d) (\x. \i\Basis. \f x\i\ *\<^sub>R i::'m) \ j" + unfolding inner_setsum_left[symmetric] integral_combine_division_topdown[OF * d] .. + also have "... \ integral UNIV (\x. \i\Basis. \f x\i\ *\<^sub>R i::'m) \ j" + apply(rule integral_subset_component_le) using assms * `j\Basis` by auto finally show ?case . qed finally show ?case . qed qed -lemma nonnegative_absolutely_integrable: fixes f::"'n::ordered_euclidean_space \ 'm::ordered_euclidean_space" - assumes "\x\s. \i f(x)$$i" "f integrable_on s" +lemma nonnegative_absolutely_integrable: + fixes f::"'n::ordered_euclidean_space \ 'm::ordered_euclidean_space" + assumes "\x\s. \i\Basis. 0 \ f(x)\i" "f integrable_on s" shows "f absolutely_integrable_on s" - unfolding absolutely_integrable_abs_eq apply rule defer - apply(rule integrable_eq[of _ f]) using assms apply-apply(subst euclidean_eq) by auto + unfolding absolutely_integrable_abs_eq + apply rule + apply (rule assms) + apply (rule integrable_eq[of _ f]) + using assms + apply (auto simp: euclidean_eq_iff[where 'a='m]) + done lemma absolutely_integrable_integrable_bound: fixes f::"'n::ordered_euclidean_space \ 'm::ordered_euclidean_space" assumes "\x\s. norm(f x) \ g x" "f integrable_on s" "g integrable_on s" @@ -5881,8 +5848,8 @@ apply(subst Inf_insert_finite) apply(rule finite_imageI[OF insert(1)]) proof(cases "k={}") case True thus ?P apply(subst if_P) defer apply(rule insert(5)[rule_format]) by auto - next case False thus ?P apply(subst if_not_P) defer - apply(rule absolutely_integrable_min[where 'n=real,unfolded Eucl_real_simps]) + next case False thus ?P apply(subst if_not_P) defer + apply (rule absolutely_integrable_min[where 'n=real, unfolded Basis_real_def, simplified]) defer apply(rule insert(3)[OF False]) using insert(5) by auto qed qed auto @@ -5897,7 +5864,7 @@ proof(cases "k={}") case True thus ?P apply(subst if_P) defer apply(rule insert(5)[rule_format]) by auto next case False thus ?P apply(subst if_not_P) defer - apply(rule absolutely_integrable_max[where 'n=real,unfolded Eucl_real_simps]) + apply (rule absolutely_integrable_max[where 'n=real, unfolded Basis_real_def, simplified]) defer apply(rule insert(3)[OF False]) using insert(5) by auto qed qed auto diff -r 46be26e02456 -r 899c9c4e4a4c src/HOL/Multivariate_Analysis/Linear_Algebra.thy --- a/src/HOL/Multivariate_Analysis/Linear_Algebra.thy Fri Dec 14 14:46:01 2012 +0100 +++ b/src/HOL/Multivariate_Analysis/Linear_Algebra.thy Fri Dec 14 15:46:01 2012 +0100 @@ -398,6 +398,53 @@ then show "h = g" by (simp add: ext) qed +text {* TODO: The following lemmas about adjoints should hold for any +Hilbert space (i.e. complete inner product space). +(see \url{http://en.wikipedia.org/wiki/Hermitian_adjoint}) +*} + +lemma adjoint_works: + fixes f:: "'n::euclidean_space \ 'm::euclidean_space" + assumes lf: "linear f" + shows "x \ adjoint f y = f x \ y" +proof - + have "\y. \w. \x. f x \ y = x \ w" + proof (intro allI exI) + fix y :: "'m" and x + let ?w = "(\i\Basis. (f i \ y) *\<^sub>R i) :: 'n" + have "f x \ y = f (\i\Basis. (x \ i) *\<^sub>R i) \ y" + by (simp add: euclidean_representation) + also have "\ = (\i\Basis. (x \ i) *\<^sub>R f i) \ y" + unfolding linear_setsum[OF lf finite_Basis] + by (simp add: linear_cmul[OF lf]) + finally show "f x \ y = x \ ?w" + by (simp add: inner_setsum_left inner_setsum_right mult_commute) + qed + then show ?thesis + unfolding adjoint_def choice_iff + by (intro someI2_ex[where Q="\f'. x \ f' y = f x \ y"]) auto +qed + +lemma adjoint_clauses: + fixes f:: "'n::euclidean_space \ 'm::euclidean_space" + assumes lf: "linear f" + shows "x \ adjoint f y = f x \ y" + and "adjoint f y \ x = y \ f x" + by (simp_all add: adjoint_works[OF lf] inner_commute) + +lemma adjoint_linear: + fixes f:: "'n::euclidean_space \ 'm::euclidean_space" + assumes lf: "linear f" + shows "linear (adjoint f)" + by (simp add: lf linear_def euclidean_eq_iff[where 'a='n] euclidean_eq_iff[where 'a='m] + adjoint_clauses[OF lf] inner_simps) + +lemma adjoint_adjoint: + fixes f:: "'n::euclidean_space \ 'm::euclidean_space" + assumes lf: "linear f" + shows "adjoint (adjoint f) = f" + by (rule adjoint_unique, simp add: adjoint_clauses [OF lf]) + subsection {* Interlude: Some properties of real sets *} lemma seq_mono_lemma: assumes "\(n::nat) \ m. (d n :: real) < e n" and "\n \ m. e n <= e m" @@ -1261,77 +1308,37 @@ subsection{* Euclidean Spaces as Typeclass*} -lemma independent_eq_inj_on: - fixes D :: nat - and f :: "nat \ 'c::real_vector" - assumes "inj_on f {.. (\a u. a < D \ (\i\{..R f i) \ f a)" -proof - - from assms have eq: "\i. i < D \ f ` {..i. inj_on f ({..i. finite (f ` {.. {DIM('a)..}" by auto - show ?thesis unfolding * image_Un basis_finite by auto -qed - -lemma (in euclidean_space) range_basis_finite[intro]: "finite (range basis)" - unfolding range_basis by auto - -lemma span_basis: "span (range basis) = (UNIV :: 'a::euclidean_space set)" -proof - - { fix x :: 'a - have "(\iR basis i) \ span (range basis :: 'a set)" - by (simp add: span_setsum span_mul span_superset) - then have "x \ span (range basis)" - by (simp only: euclidean_representation [symmetric]) - } then show ?thesis by auto -qed - -lemma basis_representation: - "\u. x = (\v\basis ` {..R (v\'a\euclidean_space))" -proof - - have "x\UNIV" by auto from this[unfolded span_basis[THEN sym]] - have "\u. (\v\basis ` {..R v) = x" - unfolding range_basis span_insert_0 apply(subst (asm) span_finite) by auto - then show ?thesis by fastforce -qed - -lemma span_basis'[simp]:"span ((basis::nat=>'a) ` {..'a) ` {.. span (basis ` {.. e \ \x$$i\ <= e" - by (metis component_le_norm order_trans) - -lemma norm_bound_component_lt: "norm (x::'a::euclidean_space) < e \ \x$$i\ < e" - by (metis component_le_norm basic_trans_rules(21)) - -lemma norm_le_l1: "norm (x::'a::euclidean_space) \ (\ix $$ i\)" - apply (subst euclidean_representation[of x]) +lemma in_span_Basis: "x \ span Basis" + unfolding span_Basis .. + +lemma Basis_le_norm: "b \ Basis \ \x \ b\ \ norm x" + by (rule order_trans [OF Cauchy_Schwarz_ineq2]) simp + +lemma norm_bound_Basis_le: "b \ Basis \ norm x \ e \ \x \ b\ \ e" + by (metis Basis_le_norm order_trans) + +lemma norm_bound_Basis_lt: "b \ Basis \ norm x < e \ \x \ b\ < e" + by (metis Basis_le_norm basic_trans_rules(21)) + +lemma norm_le_l1: "norm x \ (\b\Basis. \x \ b\)" + apply (subst euclidean_representation[of x, symmetric]) apply (rule order_trans[OF norm_setsum]) apply (auto intro!: setsum_mono) done @@ -1339,61 +1346,29 @@ lemma setsum_norm_allsubsets_bound: fixes f:: "'a \ 'n::euclidean_space" assumes fP: "finite P" and fPs: "\Q. Q \ P \ norm (setsum f Q) \ e" - shows "setsum (\x. norm (f x)) P \ 2 * real DIM('n) * e" + shows "(\x\P. norm (f x)) \ 2 * real DIM('n) * e" proof - - let ?d = "real DIM('n)" - let ?nf = "\x. norm (f x)" - let ?U = "{..x. setsum (\i. \f x $$ i\) ?U) P = setsum (\i. setsum (\x. \f x $$ i\) P) ?U" + have "(\x\P. norm (f x)) \ (\x\P. \b\Basis. \f x \ b\)" + by (rule setsum_mono) (rule norm_le_l1) + also have "(\x\P. \b\Basis. \f x \ b\) = (\b\Basis. \x\P. \f x \ b\)" by (rule setsum_commute) - have th1: "2 * ?d * e = of_nat (card ?U) * (2 * e)" by (simp add: real_of_nat_def) - have "setsum ?nf P \ setsum (\x. setsum (\i. \f x $$ i\) ?U) P" - by (rule setsum_mono) (rule norm_le_l1) - also have "\ \ 2 * ?d * e" - unfolding th0 th1 + also have "\ \ of_nat (card (Basis :: 'n set)) * (2 * e)" proof (rule setsum_bounded) - fix i assume i: "i \ ?U" - let ?Pp = "{x. x\ P \ f x $$ i \ 0}" - let ?Pn = "{x. x \ P \ f x $$ i < 0}" - have thp: "P = ?Pp \ ?Pn" by auto - have thp0: "?Pp \ ?Pn ={}" by auto - have PpP: "?Pp \ P" and PnP: "?Pn \ P" by blast+ - have Ppe:"setsum (\x. \f x $$ i\) ?Pp \ e" - using component_le_norm[of "setsum (\x. f x) ?Pp" i] fPs[OF PpP] - by(auto intro: abs_le_D1) - have Pne: "setsum (\x. \f x $$ i\) ?Pn \ e" - using component_le_norm[of "setsum (\x. - f x) ?Pn" i] fPs[OF PnP] - by(auto simp add: setsum_negf intro: abs_le_D1) - have "setsum (\x. \f x $$ i\) P = setsum (\x. \f x $$ i\) ?Pp + setsum (\x. \f x $$ i\) ?Pn" - apply (subst thp) - apply (rule setsum_Un_zero) - using fP thp0 apply auto - done - also have "\ \ 2*e" using Pne Ppe by arith - finally show "setsum (\x. \f x $$ i\) P \ 2*e" . + fix i :: 'n assume i: "i \ Basis" + have "norm (\x\P. \f x \ i\) \ + norm ((\x\P \ - {x. f x \ i < 0}. f x) \ i) + norm ((\x\P \ {x. f x \ i < 0}. f x) \ i)" + by (simp add: abs_real_def setsum_cases[OF fP] setsum_negf uminus_add_conv_diff + norm_triangle_ineq4 inner_setsum_left + del: real_norm_def) + also have "\ \ e + e" unfolding real_norm_def + by (intro add_mono norm_bound_Basis_le i fPs) auto + finally show "(\x\P. \f x \ i\) \ 2*e" by simp qed + also have "\ = 2 * real DIM('n) * e" + by (simp add: real_of_nat_def) finally show ?thesis . qed -lemma lambda_skolem': "(\ix. P i x) \ - (\x::'a. \i ?rhs") -proof - - let ?S = "{..iiB. \x. norm (f x) \ B * norm x" proof - - let ?S = "{..i. (x$$i) *\<^sub>R (basis i)) ?S))" - apply (subst euclidean_representation[of x]) - apply rule - done - also have "\ = norm (setsum (\ i. (x$$i) *\<^sub>R f (basis i)) ?S)" - using linear_setsum[OF lf fS, of ?g, unfolded o_def] linear_cmul[OF lf] by auto - finally have th0: "norm (f x) = norm (setsum (\i. (x$$i) *\<^sub>R f (basis i))?S)" . - { fix i assume i: "i \ ?S" - from component_le_norm[of x i] - have "norm ((x$$i) *\<^sub>R f (basis i :: 'a)) \ norm (f (basis i)) * norm x" + let ?g = "\b. (x \ b) *\<^sub>R f b" + have "norm (f x) = norm (f (\b\Basis. (x \ b) *\<^sub>R b))" + unfolding euclidean_representation .. + also have "\ = norm (setsum ?g Basis)" + using linear_setsum[OF lf finite_Basis, of "\b. (x \ b) *\<^sub>R b", unfolded o_def] linear_cmul[OF lf] by auto + finally have th0: "norm (f x) = norm (setsum ?g Basis)" . + { fix i :: 'a assume i: "i \ Basis" + from Basis_le_norm[OF i, of x] + have "norm (?g i) \ norm (f i) * norm x" unfolding norm_scaleR - apply (simp only: mult_commute) + apply (subst mult_commute) apply (rule mult_mono) apply (auto simp add: field_simps) done } - then have th: "\i\ ?S. norm ((x$$i) *\<^sub>R f (basis i :: 'a)) \ norm (f (basis i)) * norm x" + then have th: "\b\Basis. norm (?g b) \ norm (f b) * norm x" by metis - from setsum_norm_le[of _ "\i. (x$$i) *\<^sub>R (f (basis i))", OF th] + from setsum_norm_le[of _ ?g, OF th] have "norm (f x) \ ?B * norm x" unfolding th0 setsum_left_distrib by metis} then show ?thesis by blast qed @@ -1438,12 +1409,15 @@ let ?K = "\B\ + 1" have Kp: "?K > 0" by arith { assume C: "B < 0" - have "((\\ i. 1)::'a) \ 0" unfolding euclidean_eq[where 'a='a] - by(auto intro!:exI[where x=0]) - then have "norm ((\\ i. 1)::'a) > 0" by auto - with C have "B * norm ((\\ i. 1)::'a) < 0" + def One \ "\Basis ::'a" + then have "One \ 0" + unfolding euclidean_eq_iff[where 'a='a] + by (simp add: inner_setsum_left inner_Basis setsum_cases) + then have "norm One > 0" by auto + with C have "B * norm One < 0" by (simp add: mult_less_0_iff) - with B[rule_format, of "(\\ i. 1)::'a"] norm_ge_zero[of "f ((\\ i. 1)::'a)"] have False by simp + with B[rule_format, of One] norm_ge_zero[of "f One"] + have False by simp } then have Bp: "B \ 0" by (metis not_leE) { fix x::"'a" @@ -1492,33 +1466,27 @@ fixes h:: "'m::euclidean_space \ 'n::euclidean_space \ 'k::real_normed_vector" assumes bh: "bilinear h" shows "\B. \x y. norm (h x y) \ B * norm x * norm y" -proof - - let ?M = "{..i. (x$$i) *\<^sub>R basis i) ?M) (setsum (\i. (y$$i) *\<^sub>R basis i) ?N))" - apply(subst euclidean_representation[where 'a='m]) - apply(subst euclidean_representation[where 'a='n]) - apply rule - done - also have "\ = norm (setsum (\ (i,j). h ((x$$i) *\<^sub>R basis i) ((y$$j) *\<^sub>R basis j)) (?M \ ?N))" - unfolding bilinear_setsum[OF bh fM fN] .. - finally have th: "norm (h x y) = \" . - have "norm (h x y) \ ?B * norm x * norm y" - apply (simp add: setsum_left_distrib th) +proof (clarify intro!: exI[of _ "\i\Basis. \j\Basis. norm (h i j)"]) + fix x:: "'m" and y :: "'n" + have "norm (h x y) = norm (h (setsum (\i. (x \ i) *\<^sub>R i) Basis) (setsum (\i. (y \ i) *\<^sub>R i) Basis))" + apply(subst euclidean_representation[where 'a='m]) + apply(subst euclidean_representation[where 'a='n]) + apply rule + done + also have "\ = norm (setsum (\ (i,j). h ((x \ i) *\<^sub>R i) ((y \ j) *\<^sub>R j)) (Basis \ Basis))" + unfolding bilinear_setsum[OF bh finite_Basis finite_Basis] .. + finally have th: "norm (h x y) = \" . + show "norm (h x y) \ (\i\Basis. \j\Basis. norm (h i j)) * norm x * norm y" + apply (auto simp add: setsum_left_distrib th setsum_cartesian_product) apply (rule setsum_norm_le) - using fN fM apply simp apply (auto simp add: bilinear_rmul[OF bh] bilinear_lmul[OF bh] field_simps simp del: scaleR_scaleR) apply (rule mult_mono) - apply (auto simp add: zero_le_mult_iff component_le_norm) + apply (auto simp add: zero_le_mult_iff Basis_le_norm) apply (rule mult_mono) - apply (auto simp add: zero_le_mult_iff component_le_norm) - done } - then show ?thesis by metis + apply (auto simp add: zero_le_mult_iff Basis_le_norm) + done qed lemma bilinear_bounded_pos: @@ -1582,8 +1550,8 @@ lemma independent_bound: fixes S:: "('a::euclidean_space) set" - shows "independent S \ finite S \ card S <= DIM('a::euclidean_space)" - using independent_span_bound[of "(basis::nat=>'a) ` {.. finite S \ card S \ DIM('a::euclidean_space)" + using independent_span_bound[OF finite_Basis, of S] by auto lemma dependent_biggerset: "(finite (S::('a::euclidean_space) set) ==> card S > DIM('a)) ==> dependent S" @@ -1666,9 +1634,8 @@ text {* More lemmas about dimension. *} lemma dim_UNIV: "dim (UNIV :: ('a::euclidean_space) set) = DIM('a)" - apply (rule dim_unique[of "(basis::nat=>'a) ` {.. T \ dim S \ dim T" @@ -2256,20 +2223,9 @@ lemma linear_eq_stdbasis: assumes lf: "linear (f::'a::euclidean_space \ _)" and lg: "linear g" - and fg: "\ib\Basis. f b = g b" shows "f = g" -proof - - let ?U = "{.. (UNIV :: 'a set)" - from equalityD2[OF span_basis'[where 'a='a]] - have IU: " (UNIV :: 'a set) \ span ?I" by blast - have "f x = g x" - apply (rule linear_eq[OF lf lg IU,rule_format]) - using fg x apply auto - done - } then show ?thesis by auto -qed + using linear_eq[OF lf lg, of _ Basis] fg by auto text {* Similar results for bilinear functions. *} @@ -2303,14 +2259,9 @@ fixes f::"'a::euclidean_space \ 'b::euclidean_space \ _" assumes bf: "bilinear f" and bg: "bilinear g" - and fg: "\iji\Basis. \j\Basis. f i j = g i j" shows "f = g" -proof - - from fg have th: "\x \ (basis ` {..y\ (basis ` {..g. linear g \ g o f = id" proof - - from linear_independent_extend[OF independent_injective_image, OF independent_basis, OF lf fi] + from linear_independent_extend[OF independent_injective_image, OF independent_Basis, OF lf fi] obtain h:: "'b => 'a" where - h: "linear h" "\x \ f ` basis ` {..i f) (basis i) = id (basis i)" + h: "linear h" "\x \ f ` Basis. h x = inv f x" by blast + from h(2) have th: "\i\Basis. (h \ f) i = id i" using inv_o_cancel[OF fi, unfolded fun_eq_iff id_def o_def] by auto @@ -2336,12 +2287,12 @@ assumes lf: "linear f" and sf: "surj f" shows "\g. linear g \ f o g = id" proof - - from linear_independent_extend[OF independent_basis[where 'a='b],of "inv f"] + from linear_independent_extend[OF independent_Basis[where 'a='b],of "inv f"] obtain h:: "'b \ 'a" where - h: "linear h" "\ x\ basis ` {..x\Basis. h x = inv f x" by blast from h(2) - have th: "\ii\Basis. (f o h) i = id i" + using sf by (auto simp add: surj_iff_all) from linear_eq_stdbasis[OF linear_compose[OF h(1) lf] linear_id th] have "f o h = id" . then show ?thesis using h(1) by blast @@ -2538,18 +2489,18 @@ subsection {* Infinity norm *} -definition "infnorm (x::'a::euclidean_space) = Sup {abs(x$$i) |i. i b) |b. b \ Basis}" lemma numseg_dimindex_nonempty: "\i. i \ (UNIV :: 'n set)" by auto lemma infnorm_set_image: - "{abs((x::'a::euclidean_space)$$i) |i. ii. abs(x$$i)) ` {.. i) |i. i \ Basis} = (\i. abs(x \ i)) ` Basis" + by blast lemma infnorm_set_lemma: - shows "finite {abs((x::'a::euclidean_space)$$i) |i. i {}" + shows "finite {abs((x::'a::euclidean_space) \ i) |i. i \ Basis}" + and "{abs(x \ i) |i. i \ Basis} \ {}" unfolding infnorm_set_image by auto @@ -2557,25 +2508,26 @@ unfolding infnorm_def unfolding Sup_finite_ge_iff[ OF infnorm_set_lemma] unfolding infnorm_set_image - by auto + by (auto simp: ex_in_conv) lemma infnorm_triangle: "infnorm ((x::'a::euclidean_space) + y) \ infnorm x + infnorm y" proof - have th: "\x y (z::real). x - y <= z \ x - z <= y" by arith have th1: "\S f. f ` S = { f i| i. i \ S}" by blast have th2: "\x (y::real). abs(x + y) - abs(x) <= abs(y)" by arith - have *:"\i. i \ {.. i {}" by blast - have th1: "\S f. f ` S = { f i| i. i \ S}" by blast - show ?thesis unfolding infnorm_def - apply(subst Sup_finite_ge_iff) using Sup_finite_in[OF fS S0] - using infnorm_set_image using True apply auto - done +lemma Basis_le_infnorm: + assumes b: "b \ Basis" shows "\x \ b\ \ infnorm (x::'a::euclidean_space)" + unfolding infnorm_def +proof (subst Sup_finite_ge_iff) + let ?S = "{\x \ i\ |i. i \ Basis}" + show "finite ?S" by (rule infnorm_set_lemma) + show "?S \ {}" by auto + show "Bex ?S (op \ \x \ b\)" + using b by (auto intro!: exI[of _ b]) qed lemma infnorm_mul_lemma: "infnorm(a *\<^sub>R x) <= \a\ * infnorm x" apply (subst infnorm_def) unfolding Sup_finite_le_iff[OF infnorm_set_lemma] - unfolding infnorm_set_image ball_simps euclidean_component_scaleR abs_mult - using component_le_infnorm[of x] + unfolding infnorm_set_image ball_simps inner_scaleR abs_mult + using Basis_le_infnorm[of _ x] apply (auto intro: mult_mono) done @@ -2671,9 +2616,13 @@ lemma infnorm_le_norm: "infnorm x \ norm x" unfolding infnorm_def Sup_finite_le_iff[OF infnorm_set_lemma] unfolding infnorm_set_image ball_simps - by (metis component_le_norm) - -lemma norm_le_infnorm: "norm(x) <= sqrt(real DIM('a)) * infnorm(x::'a::euclidean_space)" + by (metis Basis_le_norm) + +lemma euclidean_inner: "inner x y = (\b\Basis. (x \ b) * (y \ b))" + by (subst (1 2) euclidean_representation[symmetric, where 'a='a]) + (simp add: inner_setsum_left inner_setsum_right setsum_cases inner_Basis ac_simps if_distrib) + +lemma norm_le_infnorm: "norm(x) <= sqrt DIM('a) * infnorm(x::'a::euclidean_space)" proof - let ?d = "DIM('a)" have "real ?d \ 0" by simp @@ -2683,13 +2632,14 @@ by (simp add: zero_le_mult_iff infnorm_pos_le) have th1: "x \ x \ (sqrt (real ?d) * infnorm x)^2" unfolding power_mult_distrib d2 - unfolding real_of_nat_def apply(subst euclidean_inner) + unfolding real_of_nat_def + apply(subst euclidean_inner) apply (subst power2_abs[symmetric]) apply (rule order_trans[OF setsum_bounded[where K="\infnorm x\\"]]) apply (auto simp add: power2_eq_square[symmetric]) apply (subst power2_abs[symmetric]) apply (rule power_mono) - unfolding infnorm_def Sup_finite_ge_iff[OF infnorm_set_lemma] + unfolding infnorm_def Sup_finite_ge_iff[OF infnorm_set_lemma] unfolding infnorm_set_image bex_simps apply (rule_tac x=i in bexI) apply auto @@ -2872,8 +2822,8 @@ subsection {* An ordering on euclidean spaces that will allow us to talk about intervals *} class ordered_euclidean_space = ord + euclidean_space + - assumes eucl_le: "x \ y \ (\i < DIM('a). x $$ i \ y $$ i)" - and eucl_less: "x < y \ (\i < DIM('a). x $$ i < y $$ i)" + assumes eucl_le: "x \ y \ (\i\Basis. x \ i \ y \ i)" + and eucl_less: "x < y \ (\i\Basis. x \ i < y \ i)" lemma eucl_less_not_refl[simp, intro!]: "\ x < (x::'a::ordered_euclidean_space)" unfolding eucl_less[where 'a='a] by auto @@ -2889,35 +2839,16 @@ lemma atLeastAtMost_singleton_euclidean[simp]: fixes a :: "'a::ordered_euclidean_space" shows "{a .. a} = {a}" - by (force simp: eucl_le[where 'a='a] euclidean_eq[where 'a='a]) - -lemma basis_real_range: "basis ` {..<1} = {1::real}" by auto - -instance real::ordered_euclidean_space - by default (auto simp add: euclidean_component_def) - -lemma Eucl_real_simps[simp]: - "(x::real) $$ 0 = x" - "(\\ i. f i) = ((f 0)::real)" - "\i. i > 0 \ x $$ i = 0" - defer apply(subst euclidean_eq) apply safe - unfolding euclidean_lambda_beta' - unfolding euclidean_component_def apply auto - done - -lemma complex_basis[simp]: - shows "basis 0 = (1::complex)" and "basis 1 = ii" and "basis (Suc 0) = ii" - unfolding basis_complex_def by auto - -lemma DIM_prod[simp]: "DIM('a \ 'b) = DIM('b::euclidean_space) + DIM('a::euclidean_space)" - (* FIXME: why this orientation? Why not "DIM('a) + DIM('b)" ? *) - unfolding dimension_prod_def by (rule add_commute) + by (force simp: eucl_le[where 'a='a] euclidean_eq_iff[where 'a='a]) + +instance real :: ordered_euclidean_space + by default (auto simp add: Basis_real_def) instantiation prod :: (ordered_euclidean_space, ordered_euclidean_space) ordered_euclidean_space begin -definition "x \ (y::('a\'b)) \ (\i'b). x $$ i \ y $$ i)" -definition "x < (y::('a\'b)) \ (\i'b). x $$ i < y $$ i)" +definition "x \ (y::('a\'b)) \ (\i\Basis. x \ i \ y \ i)" +definition "x < (y::('a\'b)) \ (\i\Basis. x \ i < y \ i)" instance by default (auto simp: less_prod_def less_eq_prod_def) diff -r 46be26e02456 -r 899c9c4e4a4c src/HOL/Multivariate_Analysis/Multivariate_Analysis.thy --- a/src/HOL/Multivariate_Analysis/Multivariate_Analysis.thy Fri Dec 14 14:46:01 2012 +0100 +++ b/src/HOL/Multivariate_Analysis/Multivariate_Analysis.thy Fri Dec 14 15:46:01 2012 +0100 @@ -1,5 +1,5 @@ theory Multivariate_Analysis -imports Fashoda Extended_Real_Limits +imports Fashoda Extended_Real_Limits Determinants begin end diff -r 46be26e02456 -r 899c9c4e4a4c src/HOL/Multivariate_Analysis/Operator_Norm.thy --- a/src/HOL/Multivariate_Analysis/Operator_Norm.thy Fri Dec 14 14:46:01 2012 +0100 +++ b/src/HOL/Multivariate_Analysis/Operator_Norm.thy Fri Dec 14 15:46:01 2012 +0100 @@ -22,8 +22,11 @@ moreover {assume H: ?lhs - have bp: "b \ 0" apply-apply(rule order_trans [OF norm_ge_zero]) - apply(rule H[rule_format, of "basis 0::'a"]) by auto + have bp: "b \ 0" + apply - + apply(rule order_trans [OF norm_ge_zero]) + apply(rule H[rule_format, of "SOME x::'a. x \ Basis"]) + by (auto intro: SOME_Basis norm_Basis) {fix x :: "'a" {assume "x = 0" then have "norm (f x) \ b * norm x" by (simp add: linear_0[OF lf] bp)} @@ -50,8 +53,8 @@ proof- { let ?S = "{norm (f x) |x. norm x = 1}" - have "norm (f (basis 0)) \ ?S" unfolding mem_Collect_eq - apply(rule_tac x="basis 0" in exI) by auto + have "norm (f (SOME i. i \ Basis)) \ ?S" + by (auto intro!: exI[of _ "SOME i. i \ Basis"] norm_Basis SOME_Basis) hence Se: "?S \ {}" by auto from linear_bounded[OF lf] have b: "\ b. ?S *<= b" unfolding norm_bound_generalize[OF lf, symmetric] by (auto simp add: setle_def) @@ -70,8 +73,8 @@ qed lemma onorm_pos_le: assumes lf: "linear (f::'n::euclidean_space \ 'm::euclidean_space)" shows "0 <= onorm f" - using order_trans[OF norm_ge_zero onorm(1)[OF lf, of "basis 0"]] - using DIM_positive[where 'a='n] by auto + using order_trans[OF norm_ge_zero onorm(1)[OF lf, of "SOME i. i \ Basis"]] + by (simp add: SOME_Basis) lemma onorm_eq_0: assumes lf: "linear (f::'a::euclidean_space \ 'b::euclidean_space)" shows "onorm f = 0 \ (\x. f x = 0)" @@ -87,7 +90,7 @@ proof- let ?f = "\x::'a. (y::'b)" have th: "{norm (?f x)| x. norm x = 1} = {norm y}" - apply safe apply(rule_tac x="basis 0" in exI) by auto + by (auto simp: SOME_Basis intro!: exI[of _ "SOME i. i \ Basis"]) show ?thesis unfolding onorm_def th apply (rule Sup_unique) by (simp_all add: setle_def) diff -r 46be26e02456 -r 899c9c4e4a4c src/HOL/Multivariate_Analysis/Path_Connected.thy --- a/src/HOL/Multivariate_Analysis/Path_Connected.thy Fri Dec 14 14:46:01 2012 +0100 +++ b/src/HOL/Multivariate_Analysis/Path_Connected.thy Fri Dec 14 15:46:01 2012 +0100 @@ -819,35 +819,40 @@ assumes "2 \ DIM('a::euclidean_space)" shows "path_connected((UNIV::'a::euclidean_space set) - {a})" proof - - let ?A = "{x::'a. \i\{.. {..\ i. a $$ i - 1) \ {x::'a. x $$ i < a $$ i}" by simp - show "path_connected {x. x $$ i < a $$ i}" unfolding euclidean_component_def - by (rule convex_imp_path_connected [OF convex_halfspace_lt]) + fix i :: 'a + assume "i \ Basis" + then show "(\i\Basis. (a \ i - 1)*\<^sub>R i) \ {x::'a. x \ i < a \ i}" by simp + show "path_connected {x. x \ i < a \ i}" + using convex_imp_path_connected [OF convex_halfspace_lt, of i "a \ i"] + by (simp add: inner_commute) qed have B: "path_connected ?B" unfolding Collect_bex_eq proof (rule path_connected_UNION) - fix i - assume "i \ {..\ i. a $$ i + 1) \ {x::'a. a $$ i < x $$ i}" by simp - show "path_connected {x. a $$ i < x $$ i}" unfolding euclidean_component_def - by (rule convex_imp_path_connected [OF convex_halfspace_gt]) + fix i :: 'a + assume "i \ Basis" + then show "(\i\Basis. (a \ i + 1) *\<^sub>R i) \ {x::'a. a \ i < x \ i}" by simp + show "path_connected {x. a \ i < x \ i}" + using convex_imp_path_connected [OF convex_halfspace_gt, of "a \ i" i] + by (simp add: inner_commute) qed - from assms have "1 < DIM('a)" by auto - then have "a + basis 0 - basis 1 \ ?A \ ?B" by auto + obtain S :: "'a set" where "S \ Basis" "card S = Suc (Suc 0)" + using ex_card[OF assms] by auto + then obtain b0 b1 :: 'a where "b0 \ Basis" "b1 \ Basis" "b0 \ b1" + unfolding card_Suc_eq by auto + then have "a + b0 - b1 \ ?A \ ?B" by (auto simp: inner_simps inner_Basis) then have "?A \ ?B \ {}" by fast with A B have "path_connected (?A \ ?B)" by (rule path_connected_Un) - also have "?A \ ?B = {x. \i\{.. a $$ i}" + also have "?A \ ?B = {x. \i\Basis. x \ i \ a \ i}" unfolding neq_iff bex_disj_distrib Collect_disj_eq .. also have "\ = {x. x \ a}" - unfolding Bex_def euclidean_eq [where 'a='a] by simp + unfolding euclidean_eq_iff [where 'a='a] by (simp add: Bex_def) also have "\ = UNIV - {a}" by auto finally show ?thesis . qed diff -r 46be26e02456 -r 899c9c4e4a4c src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy --- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Fri Dec 14 14:46:01 2012 +0100 +++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Fri Dec 14 15:46:01 2012 +0100 @@ -13,9 +13,17 @@ "~~/src/HOL/Library/Countable_Set" Linear_Algebra "~~/src/HOL/Library/Glbs" + "~~/src/HOL/Library/FuncSet" Norm_Arith begin +lemma countable_PiE: + "finite I \ (\i. i \ I \ countable (F i)) \ countable (PiE I F)" + by (induct I arbitrary: F rule: finite_induct) (auto simp: PiE_insert_eq) + +lemma countable_rat: "countable \" + unfolding Rats_def by auto + subsection {* Topological Basis *} context topological_space @@ -593,86 +601,74 @@ lemma ball_empty[intro]: "e \ 0 ==> ball x e = {}" by simp +lemma euclidean_dist_l2: + fixes x y :: "'a :: euclidean_space" + shows "dist x y = setL2 (\i. dist (x \ i) (y \ i)) Basis" + unfolding dist_norm norm_eq_sqrt_inner setL2_def + by (subst euclidean_inner) (simp add: power2_eq_square inner_diff_left) + +definition "box a b = {x. \i\Basis. a \ i < x \ i \ x \ i < b \ i}" + lemma rational_boxes: - fixes x :: "'a\ordered_euclidean_space" + fixes x :: "'a\euclidean_space" assumes "0 < e" - shows "\a b. (\i. a $$ i \ \) \ (\i. b $$ i \ \) \ x \ {a <..< b} \ {a <..< b} \ ball x e" + shows "\a b. (\i\Basis. a \ i \ \ \ b \ i \ \ ) \ x \ box a b \ box a b \ ball x e" proof - def e' \ "e / (2 * sqrt (real (DIM ('a))))" - then have e: "0 < e'" using assms by (auto intro!: divide_pos_pos) - have "\i. \y. y \ \ \ y < x $$ i \ x $$ i - y < e'" (is "\i. ?th i") + then have e: "0 < e'" using assms by (auto intro!: divide_pos_pos simp: DIM_positive) + have "\i. \y. y \ \ \ y < x \ i \ x \ i - y < e'" (is "\i. ?th i") proof - fix i from Rats_dense_in_real[of "x $$ i - e'" "x $$ i"] e - show "?th i" by auto + fix i from Rats_dense_in_real[of "x \ i - e'" "x \ i"] e show "?th i" by auto qed from choice[OF this] guess a .. note a = this - have "\i. \y. y \ \ \ x $$ i < y \ y - x $$ i < e'" (is "\i. ?th i") + have "\i. \y. y \ \ \ x \ i < y \ y - x \ i < e'" (is "\i. ?th i") proof - fix i from Rats_dense_in_real[of "x $$ i" "x $$ i + e'"] e - show "?th i" by auto + fix i from Rats_dense_in_real[of "x \ i" "x \ i + e'"] e show "?th i" by auto qed from choice[OF this] guess b .. note b = this - { fix y :: 'a assume *: "Chi a < y" "y < Chi b" - have "dist x y = sqrt (\i)" + let ?a = "\i\Basis. a i *\<^sub>R i" and ?b = "\i\Basis. b i *\<^sub>R i" + show ?thesis + proof (rule exI[of _ ?a], rule exI[of _ ?b], safe) + fix y :: 'a assume *: "y \ box ?a ?b" + have "dist x y = sqrt (\i\Basis. (dist (x \ i) (y \ i))\)" unfolding setL2_def[symmetric] by (rule euclidean_dist_l2) - also have "\ < sqrt (\i < sqrt (\(i::'a)\Basis. e^2 / real (DIM('a)))" proof (rule real_sqrt_less_mono, rule setsum_strict_mono) - fix i assume i: "i \ {.. y$$i < b i" using * i eucl_less[where 'a='a] by auto - moreover have "a i < x$$i" "x$$i - a i < e'" using a by auto - moreover have "x$$i < b i" "b i - x$$i < e'" using b by auto - ultimately have "\x$$i - y$$i\ < 2 * e'" by auto - then have "dist (x $$ i) (y $$ i) < e/sqrt (real (DIM('a)))" + fix i :: "'a" assume i: "i \ Basis" + have "a i < y\i \ y\i < b i" using * i by (auto simp: box_def) + moreover have "a i < x\i" "x\i - a i < e'" using a by auto + moreover have "x\i < b i" "b i - x\i < e'" using b by auto + ultimately have "\x\i - y\i\ < 2 * e'" by auto + then have "dist (x \ i) (y \ i) < e/sqrt (real (DIM('a)))" unfolding e'_def by (auto simp: dist_real_def) - then have "(dist (x $$ i) (y $$ i))\ < (e/sqrt (real (DIM('a))))\" + then have "(dist (x \ i) (y \ i))\ < (e/sqrt (real (DIM('a))))\" by (rule power_strict_mono) auto - then show "(dist (x $$ i) (y $$ i))\ < e\ / real DIM('a)" + then show "(dist (x \ i) (y \ i))\ < e\ / real DIM('a)" by (simp add: power_divide) qed auto - also have "\ = e" using `0 < e` by (simp add: real_eq_of_nat DIM_positive) - finally have "dist x y < e" . } - with a b show ?thesis - apply (rule_tac exI[of _ "Chi a"]) - apply (rule_tac exI[of _ "Chi b"]) - using eucl_less[where 'a='a] by auto -qed - -lemma ex_rat_list: - fixes x :: "'a\ordered_euclidean_space" - assumes "\ i. x $$ i \ \" - shows "\ r. length r = DIM('a) \ (\ i < DIM('a). of_rat (r ! i) = x $$ i)" -proof - - have "\i. \r. x $$ i = of_rat r" using assms unfolding Rats_def by blast - from choice[OF this] guess r .. - then show ?thesis by (auto intro!: exI[of _ "map r [0 ..< DIM('a)]"]) -qed - -lemma open_UNION: - fixes M :: "'a\ordered_euclidean_space set" - assumes "open M" - shows "M = UNION {(a, b) | a b. {Chi (of_rat \ op ! a) <..< Chi (of_rat \ op ! b)} \ M} - (\ (a, b). {Chi (of_rat \ op ! a) <..< Chi (of_rat \ op ! b)})" - (is "M = UNION ?idx ?box") + also have "\ = e" using `0 < e` by (simp add: real_eq_of_nat) + finally show "y \ ball x e" by (auto simp: ball_def) + qed (insert a b, auto simp: box_def) +qed + +lemma open_UNION_box: + fixes M :: "'a\euclidean_space set" + assumes "open M" + defines "a' \ \f :: 'a \ real \ real. (\(i::'a)\Basis. fst (f i) *\<^sub>R i)" + defines "b' \ \f :: 'a \ real \ real. (\(i::'a)\Basis. snd (f i) *\<^sub>R i)" + defines "I \ {f\Basis \\<^isub>E \ \ \. box (a' f) (b' f) \ M}" + shows "M = (\f\I. box (a' f) (b' f))" proof safe fix x assume "x \ M" obtain e where e: "e > 0" "ball x e \ M" - using openE[OF assms `x \ M`] by auto - then obtain a b where ab: "x \ {a <..< b}" "\i. a $$ i \ \" "\i. b $$ i \ \" "{a <..< b} \ ball x e" - using rational_boxes[OF e(1)] by blast - then obtain p q where pq: "length p = DIM ('a)" - "length q = DIM ('a)" - "\ i < DIM ('a). of_rat (p ! i) = a $$ i \ of_rat (q ! i) = b $$ i" - using ex_rat_list[OF ab(2)] ex_rat_list[OF ab(3)] by blast - hence p: "Chi (of_rat \ op ! p) = a" - using euclidean_eq[of "Chi (of_rat \ op ! p)" a] - unfolding o_def by auto - from pq have q: "Chi (of_rat \ op ! q) = b" - using euclidean_eq[of "Chi (of_rat \ op ! q)" b] - unfolding o_def by auto - have "x \ ?box (p, q)" - using p q ab by auto - thus "x \ UNION ?idx ?box" using ab e p q exI[of _ p] exI[of _ q] by auto -qed auto + using openE[OF `open M` `x \ M`] by auto + moreover then obtain a b where ab: "x \ box a b" + "\i \ Basis. a \ i \ \" "\i\Basis. b \ i \ \" "box a b \ ball x e" + using rational_boxes[OF e(1)] by metis + ultimately show "x \ (\f\I. box (a' f) (b' f))" + by (intro UN_I[of "\i\Basis. (a \ i, b \ i)"]) + (auto simp: euclidean_representation I_def a'_def b'_def) +qed (auto simp: I_def) subsection{* Connectedness *} @@ -1156,14 +1152,10 @@ "eventually P (at a) \ (\d>0. \x. 0 < dist x a \ dist x a < d \ P x)" unfolding eventually_at dist_nz by auto -lemma eventually_within: "eventually P (at a within S) \ +lemma eventually_within: (* FIXME: this replaces Limits.eventually_within *) + "eventually P (at a within S) \ (\d>0. \x\S. 0 < dist x a \ dist x a < d \ P x)" -unfolding eventually_within eventually_at dist_nz by auto - -lemma eventually_within_le: "eventually P (at a within S) \ - (\d>0. \x\S. 0 < dist x a \ dist x a <= d \ P x)" (is "?lhs = ?rhs") -unfolding eventually_within -by auto (metis dense order_le_less_trans) + by (rule eventually_within_less) lemma eventually_happens: "eventually P net ==> trivial_limit net \ (\x. P x)" unfolding trivial_limit_def @@ -1721,7 +1713,7 @@ assume "\ (\y\A. dist y x < e)" hence "infdist x A \ e" using `a \ A` unfolding infdist_def - by (force intro: Inf_greatest simp: dist_commute) + by (force simp: dist_commute) with x `0 < e` show False by auto qed qed @@ -2374,56 +2366,41 @@ by auto qed -lemma bounded_component: "bounded s \ bounded ((\x. x $$ i) ` s)" - apply (erule bounded_linear_image) - apply (rule bounded_linear_euclidean_component) - done - lemma compact_lemma: fixes f :: "nat \ 'a::euclidean_space" assumes "bounded s" and "\n. f n \ s" - shows "\d. \l::'a. \ r. subseq r \ - (\e>0. eventually (\n. \i\d. dist (f (r n) $$ i) (l $$ i) < e) sequentially)" -proof - fix d'::"nat set" def d \ "d' \ {..{..l::'a. \r. subseq r \ - (\e>0. eventually (\n. \i\d. dist (f (r n) $$ i) (l $$ i) < e) sequentially)" + shows "\d\Basis. \l::'a. \ r. subseq r \ + (\e>0. eventually (\n. \i\d. dist (f (r n) \ i) (l \ i) < e) sequentially)" +proof safe + fix d :: "'a set" assume d: "d \ Basis" + with finite_Basis have "finite d" by (blast intro: finite_subset) + from this d show "\l::'a. \r. subseq r \ + (\e>0. eventually (\n. \i\d. dist (f (r n) \ i) (l \ i) < e) sequentially)" proof(induct d) case empty thus ?case unfolding subseq_def by auto - next case (insert k d) have k[intro]:"kx. x $$ k) ` s)" using `bounded s` by (rule bounded_component) + next case (insert k d) have k[intro]:"k\Basis" using insert by auto + have s': "bounded ((\x. x \ k) ` s)" using `bounded s` + by (auto intro!: bounded_linear_image bounded_linear_inner_left) obtain l1::"'a" and r1 where r1:"subseq r1" and - lr1:"\e>0. eventually (\n. \i\d. dist (f (r1 n) $$ i) (l1 $$ i) < e) sequentially" + lr1:"\e>0. eventually (\n. \i\d. dist (f (r1 n) \ i) (l1 \ i) < e) sequentially" using insert(3) using insert(4) by auto - have f': "\n. f (r1 n) $$ k \ (\x. x $$ k) ` s" using `\n. f n \ s` by simp - obtain l2 r2 where r2:"subseq r2" and lr2:"((\i. f (r1 (r2 i)) $$ k) ---> l2) sequentially" + have f': "\n. f (r1 n) \ k \ (\x. x \ k) ` s" using `\n. f n \ s` by simp + obtain l2 r2 where r2:"subseq r2" and lr2:"((\i. f (r1 (r2 i)) \ k) ---> l2) sequentially" using bounded_imp_convergent_subsequence[OF s' f'] unfolding o_def by auto def r \ "r1 \ r2" have r:"subseq r" using r1 and r2 unfolding r_def o_def subseq_def by auto moreover - def l \ "(\\ i. if i = k then l2 else l1$$i)::'a" + def l \ "(\i\Basis. (if i = k then l2 else l1\i) *\<^sub>R i)::'a" { fix e::real assume "e>0" - from lr1 `e>0` have N1:"eventually (\n. \i\d. dist (f (r1 n) $$ i) (l1 $$ i) < e) sequentially" by blast - from lr2 `e>0` have N2:"eventually (\n. dist (f (r1 (r2 n)) $$ k) l2 < e) sequentially" by (rule tendstoD) - from r2 N1 have N1': "eventually (\n. \i\d. dist (f (r1 (r2 n)) $$ i) (l1 $$ i) < e) sequentially" + from lr1 `e>0` have N1:"eventually (\n. \i\d. dist (f (r1 n) \ i) (l1 \ i) < e) sequentially" by blast + from lr2 `e>0` have N2:"eventually (\n. dist (f (r1 (r2 n)) \ k) l2 < e) sequentially" by (rule tendstoD) + from r2 N1 have N1': "eventually (\n. \i\d. dist (f (r1 (r2 n)) \ i) (l1 \ i) < e) sequentially" by (rule eventually_subseq) - have "eventually (\n. \i\(insert k d). dist (f (r n) $$ i) (l $$ i) < e) sequentially" - using N1' N2 apply(rule eventually_elim2) unfolding l_def r_def o_def - using insert.prems by auto + have "eventually (\n. \i\(insert k d). dist (f (r n) \ i) (l \ i) < e) sequentially" + using N1' N2 + by eventually_elim (insert insert.prems, auto simp: l_def r_def o_def) } ultimately show ?case by auto qed - thus "\l::'a. \r. subseq r \ - (\e>0. eventually (\n. \i\d'. dist (f (r n) $$ i) (l $$ i) < e) sequentially)" - apply safe apply(rule_tac x=l in exI,rule_tac x=r in exI) apply safe - apply(erule_tac x=e in allE) unfolding d_def eventually_sequentially apply safe - apply(rule_tac x=N in exI) apply safe apply(erule_tac x=n in allE,safe) - apply(erule_tac x=i in ballE) - proof- fix i and r::"nat=>nat" and n::nat and e::real and l::'a - assume "i\d'" "i \ d' \ {..0" - hence *:"i\DIM('a)" by auto - thus "dist (f (r n) $$ i) (l $$ i) < e" using e by auto - qed qed instance euclidean_space \ heine_borel @@ -2431,22 +2408,20 @@ fix s :: "'a set" and f :: "nat \ 'a" assume s: "bounded s" and f: "\n. f n \ s" then obtain l::'a and r where r: "subseq r" - and l: "\e>0. eventually (\n. \i\UNIV. dist (f (r n) $$ i) (l $$ i) < e) sequentially" + and l: "\e>0. eventually (\n. \i\Basis. dist (f (r n) \ i) (l \ i) < e) sequentially" using compact_lemma [OF s f] by blast - let ?d = "{..0" - hence "0 < e / (real_of_nat (card ?d))" - using DIM_positive using divide_pos_pos[of e, of "real_of_nat (card ?d)"] by auto - with l have "eventually (\n. \i. dist (f (r n) $$ i) (l $$ i) < e / (real_of_nat (card ?d))) sequentially" + hence "0 < e / real_of_nat DIM('a)" by (auto intro!: divide_pos_pos DIM_positive) + with l have "eventually (\n. \i\Basis. dist (f (r n) \ i) (l \ i) < e / (real_of_nat DIM('a))) sequentially" by simp moreover - { fix n assume n: "\i. dist (f (r n) $$ i) (l $$ i) < e / (real_of_nat (card ?d))" - have "dist (f (r n)) l \ (\i\?d. dist (f (r n) $$ i) (l $$ i))" + { fix n assume n: "\i\Basis. dist (f (r n) \ i) (l \ i) < e / (real_of_nat DIM('a))" + have "dist (f (r n)) l \ (\i\Basis. dist (f (r n) \ i) (l \ i))" apply(subst euclidean_dist_l2) using zero_le_dist by (rule setL2_le_setsum) - also have "\ < (\i\?d. e / (real_of_nat (card ?d)))" + also have "\ < (\i\(Basis::'a set). e / (real_of_nat DIM('a)))" apply(rule setsum_strict_mono) using n by auto - finally have "dist (f (r n)) l < e" unfolding setsum_constant - using DIM_positive[where 'a='a] by auto + finally have "dist (f (r n)) l < e" + by auto } ultimately have "eventually (\n. dist (f (r n)) l < e) sequentially" by (rule eventually_elim1) @@ -3885,10 +3860,6 @@ shows "continuous F (\x. inner (f x) (g x))" using assms unfolding continuous_def by (rule tendsto_inner) -lemma continuous_euclidean_component: - shows "continuous F f \ continuous F (\x. f x $$ i)" - unfolding continuous_def by (rule tendsto_euclidean_component) - lemma continuous_inverse: fixes f :: "'a::t2_space \ 'b::real_normed_div_algebra" assumes "continuous F f" and "f (netlimit F) \ 0" @@ -3909,10 +3880,8 @@ lemmas continuous_intros = continuous_at_id continuous_within_id continuous_const continuous_dist continuous_norm continuous_infnorm - continuous_add continuous_minus continuous_diff - continuous_scaleR continuous_mult - continuous_inner continuous_euclidean_component - continuous_at_inverse continuous_at_within_inverse + continuous_add continuous_minus continuous_diff continuous_scaleR continuous_mult + continuous_inner continuous_at_inverse continuous_at_within_inverse subsubsection {* Structural rules for setwise continuity *} @@ -3976,11 +3945,6 @@ using bounded_bilinear_inner assms by (rule bounded_bilinear.continuous_on) -lemma continuous_on_euclidean_component: - "continuous_on s f \ continuous_on s (\x. f x $$ i)" - using bounded_linear_euclidean_component - by (rule bounded_linear.continuous_on) - lemma continuous_on_inverse: fixes f :: "'a::topological_space \ 'b::real_normed_div_algebra" assumes "continuous_on s f" and "\x\s. f x \ 0" @@ -4110,7 +4074,7 @@ continuous_on_compose continuous_on_norm continuous_on_infnorm continuous_on_add continuous_on_minus continuous_on_diff continuous_on_scaleR continuous_on_mult continuous_on_inverse - continuous_on_inner continuous_on_euclidean_component + continuous_on_inner uniformly_continuous_on_id uniformly_continuous_on_const uniformly_continuous_on_dist uniformly_continuous_on_norm uniformly_continuous_on_compose uniformly_continuous_on_add @@ -5070,65 +5034,64 @@ subsection {* Intervals *} lemma interval: fixes a :: "'a::ordered_euclidean_space" shows - "{a <..< b} = {x::'a. \i x$$i < b$$i}" and - "{a .. b} = {x::'a. \i x$$i \ x$$i \ b$$i}" + "{a <..< b} = {x::'a. \i\Basis. a\i < x\i \ x\i < b\i}" and + "{a .. b} = {x::'a. \i\Basis. a\i \ x\i \ x\i \ b\i}" by(auto simp add:set_eq_iff eucl_le[where 'a='a] eucl_less[where 'a='a]) lemma mem_interval: fixes a :: "'a::ordered_euclidean_space" shows - "x \ {a<.. (\i x$$i < b$$i)" - "x \ {a .. b} \ (\i x$$i \ x$$i \ b$$i)" + "x \ {a<.. (\i\Basis. a\i < x\i \ x\i < b\i)" + "x \ {a .. b} \ (\i\Basis. a\i \ x\i \ x\i \ b\i)" using interval[of a b] by(auto simp add: set_eq_iff eucl_le[where 'a='a] eucl_less[where 'a='a]) lemma interval_eq_empty: fixes a :: "'a::ordered_euclidean_space" shows - "({a <..< b} = {} \ (\i a$$i))" (is ?th1) and - "({a .. b} = {} \ (\i (\i\Basis. b\i \ a\i))" (is ?th1) and + "({a .. b} = {} \ (\i\Basis. b\i < a\i))" (is ?th2) proof- - { fix i x assume i:"i a$$i" and x:"x\{a <..< b}" - hence "a $$ i < x $$ i \ x $$ i < b $$ i" unfolding mem_interval by auto - hence "a$$i < b$$i" by auto + { fix i x assume i:"i\Basis" and as:"b\i \ a\i" and x:"x\{a <..< b}" + hence "a \ i < x \ i \ x \ i < b \ i" unfolding mem_interval by auto + hence "a\i < b\i" by auto hence False using as by auto } moreover - { assume as:"\i (b$$i \ a$$i)" + { assume as:"\i\Basis. \ (b\i \ a\i)" let ?x = "(1/2) *\<^sub>R (a + b)" - { fix i assume i:"iR (a+b)) $$ i" "((1/2) *\<^sub>R (a+b)) $$ i < b$$i" - unfolding euclidean_simps by auto } + { fix i :: 'a assume i:"i\Basis" + have "a\i < b\i" using as[THEN bspec[where x=i]] i by auto + hence "a\i < ((1/2) *\<^sub>R (a+b)) \ i" "((1/2) *\<^sub>R (a+b)) \ i < b\i" + by (auto simp: inner_add_left) } hence "{a <..< b} \ {}" using mem_interval(1)[of "?x" a b] by auto } ultimately show ?th1 by blast - { fix i x assume i:"i{a .. b}" - hence "a $$ i \ x $$ i \ x $$ i \ b $$ i" unfolding mem_interval by auto - hence "a$$i \ b$$i" by auto + { fix i x assume i:"i\Basis" and as:"b\i < a\i" and x:"x\{a .. b}" + hence "a \ i \ x \ i \ x \ i \ b \ i" unfolding mem_interval by auto + hence "a\i \ b\i" by auto hence False using as by auto } moreover - { assume as:"\i (b$$i < a$$i)" + { assume as:"\i\Basis. \ (b\i < a\i)" let ?x = "(1/2) *\<^sub>R (a + b)" - { fix i assume i:"i b$$i" using as[THEN spec[where x=i]] by auto - hence "a$$i \ ((1/2) *\<^sub>R (a+b)) $$ i" "((1/2) *\<^sub>R (a+b)) $$ i \ b$$i" - unfolding euclidean_simps by auto } + { fix i :: 'a assume i:"i\Basis" + have "a\i \ b\i" using as[THEN bspec[where x=i]] i by auto + hence "a\i \ ((1/2) *\<^sub>R (a+b)) \ i" "((1/2) *\<^sub>R (a+b)) \ i \ b\i" + by (auto simp: inner_add_left) } hence "{a .. b} \ {}" using mem_interval(2)[of "?x" a b] by auto } ultimately show ?th2 by blast qed lemma interval_ne_empty: fixes a :: "'a::ordered_euclidean_space" shows - "{a .. b} \ {} \ (\i b$$i)" and - "{a <..< b} \ {} \ (\i {} \ (\i\Basis. a\i \ b\i)" and + "{a <..< b} \ {} \ (\i\Basis. a\i < b\i)" unfolding interval_eq_empty[of a b] by fastforce+ lemma interval_sing: fixes a :: "'a::ordered_euclidean_space" shows "{a .. a} = {a}" and "{a<..i c$$i \ d$$i \ b$$i) \ {c .. d} \ {a .. b}" and - "(\i d$$i < b$$i) \ {c .. d} \ {a<..i c$$i \ d$$i \ b$$i) \ {c<.. {a .. b}" and - "(\i c$$i \ d$$i \ b$$i) \ {c<.. {a<..i\Basis. a\i \ c\i \ d\i \ b\i) \ {c .. d} \ {a .. b}" and + "(\i\Basis. a\i < c\i \ d\i < b\i) \ {c .. d} \ {a<..i\Basis. a\i \ c\i \ d\i \ b\i) \ {c<.. {a .. b}" and + "(\i\Basis. a\i \ c\i \ d\i \ b\i) \ {c<.. {a<.. {a .. b} \ (\i d$$i) --> (\i c$$i \ d$$i \ b$$i)" (is ?th1) and - "{c .. d} \ {a<.. (\i d$$i) --> (\i d$$i < b$$i)" (is ?th2) and - "{c<.. {a .. b} \ (\i (\i c$$i \ d$$i \ b$$i)" (is ?th3) and - "{c<.. {a<.. (\i (\i c$$i \ d$$i \ b$$i)" (is ?th4) + "{c .. d} \ {a .. b} \ (\i\Basis. c\i \ d\i) --> (\i\Basis. a\i \ c\i \ d\i \ b\i)" (is ?th1) and + "{c .. d} \ {a<.. (\i\Basis. c\i \ d\i) --> (\i\Basis. a\i < c\i \ d\i < b\i)" (is ?th2) and + "{c<.. {a .. b} \ (\i\Basis. c\i < d\i) --> (\i\Basis. a\i \ c\i \ d\i \ b\i)" (is ?th3) and + "{c<.. {a<.. (\i\Basis. c\i < d\i) --> (\i\Basis. a\i \ c\i \ d\i \ b\i)" (is ?th4) proof- show ?th1 unfolding subset_eq and Ball_def and mem_interval by (auto intro: order_trans) show ?th2 unfolding subset_eq and Ball_def and mem_interval by (auto intro: le_less_trans less_le_trans order_trans less_imp_le) - { assume as: "{c<.. {a .. b}" "\i {a .. b}" "\i\Basis. c\i < d\i" hence "{c<.. {}" unfolding interval_eq_empty by auto - fix i assume i:"iBasis" (** TODO combine the following two parts as done in the HOL_light version. **) - { let ?x = "(\\ j. (if j=i then ((min (a$$j) (d$$j))+c$$j)/2 else (c$$j+d$$j)/2))::'a" - assume as2: "a$$i > c$$i" - { fix j assume j:"j ?x $$ j < d $$ j" - apply(cases "j=i") using as(2)[THEN spec[where x=j]] i + { let ?x = "(\j\Basis. (if j=i then ((min (a\j) (d\j))+c\j)/2 else (c\j+d\j)/2) *\<^sub>R j)::'a" + assume as2: "a\i > c\i" + { fix j :: 'a assume j:"j\Basis" + hence "c \ j < ?x \ j \ ?x \ j < d \ j" + apply(cases "j=i") using as(2)[THEN bspec[where x=j]] i by (auto simp add: as2) } hence "?x\{c<..{a .. b}" - unfolding mem_interval apply auto apply(rule_tac x=i in exI) - using as(2)[THEN spec[where x=i]] and as2 i + unfolding mem_interval apply auto apply(rule_tac x=i in bexI) + using as(2)[THEN bspec[where x=i]] and as2 i by auto ultimately have False using as by auto } - hence "a$$i \ c$$i" by(rule ccontr)auto + hence "a\i \ c\i" by(rule ccontr)auto moreover - { let ?x = "(\\ j. (if j=i then ((max (b$$j) (c$$j))+d$$j)/2 else (c$$j+d$$j)/2))::'a" - assume as2: "b$$i < d$$i" - { fix j assume "j ?x $$ j \ ?x $$ j > c $$ j" - apply(cases "j=i") using as(2)[THEN spec[where x=j]] - by (auto simp add: as2) } + { let ?x = "(\j\Basis. (if j=i then ((max (b\j) (c\j))+d\j)/2 else (c\j+d\j)/2) *\<^sub>R j)::'a" + assume as2: "b\i < d\i" + { fix j :: 'a assume "j\Basis" + hence "d \ j > ?x \ j \ ?x \ j > c \ j" + apply(cases "j=i") using as(2)[THEN bspec[where x=j]] + by (auto simp add: as2) } hence "?x\{c<..{a .. b}" - unfolding mem_interval apply auto apply(rule_tac x=i in exI) - using as(2)[THEN spec[where x=i]] and as2 using i + unfolding mem_interval apply auto apply(rule_tac x=i in bexI) + using as(2)[THEN bspec[where x=i]] and as2 using i by auto ultimately have False using as by auto } - hence "b$$i \ d$$i" by(rule ccontr)auto + hence "b\i \ d\i" by(rule ccontr)auto ultimately - have "a$$i \ c$$i \ d$$i \ b$$i" by auto + have "a\i \ c\i \ d\i \ b\i" by auto } note part1 = this - show ?th3 unfolding subset_eq and Ball_def and mem_interval - apply(rule,rule,rule,rule) apply(rule part1) unfolding subset_eq and Ball_def and mem_interval - prefer 4 apply auto by(erule_tac x=i in allE,erule_tac x=i in allE,fastforce)+ - { assume as:"{c<.. {a<..i {a<..i\Basis. c\i < d\i" + fix i :: 'a assume i:"i\Basis" from as(1) have "{c<.. {a..b}" using interval_open_subset_closed[of a b] by auto - hence "a$$i \ c$$i \ d$$i \ b$$i" using part1 and as(2) using i by auto } note * = this + hence "a\i \ c\i \ d\i \ b\i" using part1 and as(2) using i by auto } note * = this show ?th4 unfolding subset_eq and Ball_def and mem_interval apply(rule,rule,rule,rule) apply(rule *) unfolding subset_eq and Ball_def and mem_interval prefer 4 - apply auto by(erule_tac x=i in allE, simp)+ -qed - -lemma disjoint_interval: fixes a::"'a::ordered_euclidean_space" shows - "{a .. b} \ {c .. d} = {} \ (\i d$$i < c$$i \ b$$i < c$$i \ d$$i < a$$i))" (is ?th1) and - "{a .. b} \ {c<.. (\i d$$i \ c$$i \ b$$i \ c$$i \ d$$i \ a$$i))" (is ?th2) and - "{a<.. {c .. d} = {} \ (\i a$$i \ d$$i < c$$i \ b$$i \ c$$i \ d$$i \ a$$i))" (is ?th3) and - "{a<.. {c<.. (\i a$$i \ d$$i \ c$$i \ b$$i \ c$$i \ d$$i \ a$$i))" (is ?th4) -proof- - let ?z = "(\\ i. ((max (a$$i) (c$$i)) + (min (b$$i) (d$$i))) / 2)::'a" - note * = set_eq_iff Int_iff empty_iff mem_interval all_conj_distrib[THEN sym] eq_False - show ?th1 unfolding * apply safe apply(erule_tac x="?z" in allE) - unfolding not_all apply(erule exE,rule_tac x=x in exI) apply(erule_tac[2-] x=i in allE) by auto - show ?th2 unfolding * apply safe apply(erule_tac x="?z" in allE) - unfolding not_all apply(erule exE,rule_tac x=x in exI) apply(erule_tac[2-] x=i in allE) by auto - show ?th3 unfolding * apply safe apply(erule_tac x="?z" in allE) - unfolding not_all apply(erule exE,rule_tac x=x in exI) apply(erule_tac[2-] x=i in allE) by auto - show ?th4 unfolding * apply safe apply(erule_tac x="?z" in allE) - unfolding not_all apply(erule exE,rule_tac x=x in exI) apply(erule_tac[2-] x=i in allE) by auto + apply auto by(erule_tac x=xa in allE, simp)+ qed lemma inter_interval: fixes a :: "'a::ordered_euclidean_space" shows - "{a .. b} \ {c .. d} = {(\\ i. max (a$$i) (c$$i)) .. (\\ i. min (b$$i) (d$$i))}" - unfolding set_eq_iff and Int_iff and mem_interval - by auto + "{a .. b} \ {c .. d} = {(\i\Basis. max (a\i) (c\i) *\<^sub>R i) .. (\i\Basis. min (b\i) (d\i) *\<^sub>R i)}" + unfolding set_eq_iff and Int_iff and mem_interval by auto + +lemma disjoint_interval: fixes a::"'a::ordered_euclidean_space" shows + "{a .. b} \ {c .. d} = {} \ (\i\Basis. (b\i < a\i \ d\i < c\i \ b\i < c\i \ d\i < a\i))" (is ?th1) and + "{a .. b} \ {c<.. (\i\Basis. (b\i < a\i \ d\i \ c\i \ b\i \ c\i \ d\i \ a\i))" (is ?th2) and + "{a<.. {c .. d} = {} \ (\i\Basis. (b\i \ a\i \ d\i < c\i \ b\i \ c\i \ d\i \ a\i))" (is ?th3) and + "{a<.. {c<.. (\i\Basis. (b\i \ a\i \ d\i \ c\i \ b\i \ c\i \ d\i \ a\i))" (is ?th4) +proof- + let ?z = "(\i\Basis. (((max (a\i) (c\i)) + (min (b\i) (d\i))) / 2) *\<^sub>R i)::'a" + have **: "\P Q. (\i :: 'a. i \ Basis \ Q ?z i \ P i) \ + (\i x :: 'a. i \ Basis \ P i \ Q x i) \ (\x. \i\Basis. Q x i) \ (\i\Basis. P i)" + by blast + note * = set_eq_iff Int_iff empty_iff mem_interval ball_conj_distrib[symmetric] eq_False ball_simps(10) + show ?th1 unfolding * by (intro **) auto + show ?th2 unfolding * by (intro **) auto + show ?th3 unfolding * by (intro **) auto + show ?th4 unfolding * by (intro **) auto +qed (* Moved interval_open_subset_closed a bit upwards *) lemma open_interval[intro]: fixes a b :: "'a::ordered_euclidean_space" shows "open {a<..ix. x$$i) -` {a$$i<..i\Basis. (\x. x\i) -` {a\i<..i})" by (intro open_INT finite_lessThan ballI continuous_open_vimage allI - linear_continuous_at bounded_linear_euclidean_component - open_real_greaterThanLessThan) - also have "(\ix. x$$i) -` {a$$i<..i\Basis. (\x. x\i) -` {a\i<..i}) = {a<..ix. x$$i) -` {a$$i .. b$$i})" + have "closed (\i\Basis. (\x. x\i) -` {a\i .. b\i})" by (intro closed_INT ballI continuous_closed_vimage allI - linear_continuous_at bounded_linear_euclidean_component - closed_real_atLeastAtMost) - also have "(\ix. x$$i) -` {a$$i .. b$$i}) = {a .. b}" + linear_continuous_at closed_real_atLeastAtMost finite_Basis bounded_linear_inner_left) + also have "(\i\Basis. (\x. x\i) -` {a\i .. b\i}) = {a .. b}" by (auto simp add: eucl_le [where 'a='a]) finally show "closed {a .. b}" . qed @@ -5253,29 +5217,29 @@ { fix x assume "x \ interior {a..b}" then obtain s where s:"open s" "x \ s" "s \ {a..b}" .. then obtain e where "e>0" and e:"\x'. dist x' x < e \ x' \ {a..b}" unfolding open_dist and subset_eq by auto - { fix i assume i:"iR basis i) x < e" - "dist (x + (e / 2) *\<^sub>R basis i) x < e" + { fix i :: 'a assume i:"i\Basis" + have "dist (x - (e / 2) *\<^sub>R i) x < e" + "dist (x + (e / 2) *\<^sub>R i) x < e" unfolding dist_norm apply auto - unfolding norm_minus_cancel using norm_basis and `e>0` by auto - hence "a $$ i \ (x - (e / 2) *\<^sub>R basis i) $$ i" - "(x + (e / 2) *\<^sub>R basis i) $$ i \ b $$ i" - using e[THEN spec[where x="x - (e/2) *\<^sub>R basis i"]] - and e[THEN spec[where x="x + (e/2) *\<^sub>R basis i"]] + unfolding norm_minus_cancel using norm_Basis[OF i] `e>0` by auto + hence "a \ i \ (x - (e / 2) *\<^sub>R i) \ i" + "(x + (e / 2) *\<^sub>R i) \ i \ b \ i" + using e[THEN spec[where x="x - (e/2) *\<^sub>R i"]] + and e[THEN spec[where x="x + (e/2) *\<^sub>R i"]] unfolding mem_interval using i by blast+ - hence "a $$ i < x $$ i" and "x $$ i < b $$ i" unfolding euclidean_simps - unfolding basis_component using `e>0` i by auto } + hence "a \ i < x \ i" and "x \ i < b \ i" + using `e>0` i by (auto simp: inner_diff_left inner_Basis inner_add_left) } hence "x \ {a<.. ?R" .. qed lemma bounded_closed_interval: fixes a :: "'a::ordered_euclidean_space" shows "bounded {a .. b}" proof- - let ?b = "\ia$$i\ + \b$$i\" - { fix x::"'a" assume x:"\i x $$ i \ x $$ i \ b $$ i" - { fix i assume "ix$$i\ \ \a$$i\ + \b$$i\" using x[THEN spec[where x=i]] by auto } - hence "(\ix $$ i\) \ ?b" apply-apply(rule setsum_mono) by auto + let ?b = "\i\Basis. \a\i\ + \b\i\" + { fix x::"'a" assume x:"\i\Basis. a \ i \ x \ i \ x \ i \ b \ i" + { fix i :: 'a assume "i\Basis" + hence "\x\i\ \ \a\i\ + \b\i\" using x[THEN bspec[where x=i]] by auto } + hence "(\i\Basis. \x \ i\) \ ?b" apply-apply(rule setsum_mono) by auto hence "norm x \ ?b" using norm_le_l1[of x] by auto } thus ?thesis unfolding interval and bounded_iff by auto qed @@ -5298,10 +5262,9 @@ lemma open_interval_midpoint: fixes a :: "'a::ordered_euclidean_space" assumes "{a<.. {}" shows "((1/2) *\<^sub>R (a + b)) \ {a<..R (a + b)) $$ i \ ((1 / 2) *\<^sub>R (a + b)) $$ i < b $$ i" - using assms[unfolded interval_ne_empty, THEN spec[where x=i]] - unfolding euclidean_simps by auto } + { fix i :: 'a assume "i\Basis" + hence "a \ i < ((1 / 2) *\<^sub>R (a + b)) \ i \ ((1 / 2) *\<^sub>R (a + b)) \ i < b \ i" + using assms[unfolded interval_ne_empty, THEN bspec[where x=i]] by (auto simp: inner_add_left) } thus ?thesis unfolding mem_interval by auto qed @@ -5309,23 +5272,23 @@ assumes x:"x \ {a<.. {a .. b}" and e:"0 < e" "e \ 1" shows "(e *\<^sub>R x + (1 - e) *\<^sub>R y) \ {a<.. < e * x $$ i + (1 - e) * y $$ i" apply(rule add_less_le_mono) + { fix i :: 'a assume i:"i\Basis" + have "a \ i = e * (a \ i) + (1 - e) * (a \ i)" unfolding left_diff_distrib by simp + also have "\ < e * (x \ i) + (1 - e) * (y \ i)" apply(rule add_less_le_mono) using e unfolding mult_less_cancel_left and mult_le_cancel_left apply simp_all using x unfolding mem_interval using i apply simp using y unfolding mem_interval using i apply simp done - finally have "a $$ i < (e *\<^sub>R x + (1 - e) *\<^sub>R y) $$ i" unfolding euclidean_simps by auto + finally have "a \ i < (e *\<^sub>R x + (1 - e) *\<^sub>R y) \ i" unfolding inner_simps by auto moreover { - have "b $$ i = e * b$$i + (1 - e) * b$$i" unfolding left_diff_distrib by simp - also have "\ > e * x $$ i + (1 - e) * y $$ i" apply(rule add_less_le_mono) + have "b \ i = e * (b\i) + (1 - e) * (b\i)" unfolding left_diff_distrib by simp + also have "\ > e * (x \ i) + (1 - e) * (y \ i)" apply(rule add_less_le_mono) using e unfolding mult_less_cancel_left and mult_le_cancel_left apply simp_all using x unfolding mem_interval using i apply simp using y unfolding mem_interval using i apply simp done - finally have "(e *\<^sub>R x + (1 - e) *\<^sub>R y) $$ i < b $$ i" unfolding euclidean_simps by auto - } ultimately have "a $$ i < (e *\<^sub>R x + (1 - e) *\<^sub>R y) $$ i \ (e *\<^sub>R x + (1 - e) *\<^sub>R y) $$ i < b $$ i" by auto } + finally have "(e *\<^sub>R x + (1 - e) *\<^sub>R y) \ i < b \ i" unfolding inner_simps by auto + } ultimately have "a \ i < (e *\<^sub>R x + (1 - e) *\<^sub>R y) \ i \ (e *\<^sub>R x + (1 - e) *\<^sub>R y) \ i < b \ i" by auto } thus ?thesis unfolding mem_interval by auto qed @@ -5365,11 +5328,11 @@ assumes "bounded s" shows "\a. s \ {-a<..0" and b:"\x\s. norm x \ b" using assms[unfolded bounded_pos] by auto - def a \ "(\\ i. b+1)::'a" + def a \ "(\i\Basis. (b + 1) *\<^sub>R i)::'a" { fix x assume "x\s" - fix i assume i:"is`] - and component_le_norm[of x i] unfolding euclidean_simps and a_def by auto } + fix i :: 'a assume i:"i\Basis" + hence "(-a)\i < x\i" and "x\i < a\i" using b[THEN bspec[where x=x], OF `x\s`] + and Basis_le_norm[OF i, of x] unfolding inner_simps and a_def by auto } thus ?thesis by(auto intro: exI[where x=a] simp add: eucl_less[where 'a='a]) qed @@ -5413,59 +5376,61 @@ (* Some stuff for half-infinite intervals too; FIXME: notation? *) lemma closed_interval_left: fixes b::"'a::euclidean_space" - shows "closed {x::'a. \i b$$i}" + shows "closed {x::'a. \i\Basis. x\i \ b\i}" proof- - { fix i assume i:"ie>0. \x'\{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"]] using i by auto - hence False using component_le_norm[of "y - x" i] unfolding dist_norm euclidean_simps using i - by auto } - hence "x$$i \ b$$i" by(rule ccontr)auto } + { fix i :: 'a assume i:"i\Basis" + fix x::"'a" assume x:"\e>0. \x'\{x. \i\Basis. 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"]] using i by auto + hence False using Basis_le_norm[OF i, of "y - x"] unfolding dist_norm inner_simps using i + by auto } + hence "x\i \ b\i" by(rule ccontr)auto } thus ?thesis unfolding closed_limpt unfolding islimpt_approachable by blast qed lemma closed_interval_right: fixes a::"'a::euclidean_space" - shows "closed {x::'a. \i x$$i}" + shows "closed {x::'a. \i\Basis. a\i \ x\i}" proof- - { fix i assume i:"ie>0. \x'\{x. \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"]] i by auto - hence False using component_le_norm[of "y - x" i] unfolding dist_norm and euclidean_simps by auto } - hence "a$$i \ x$$i" by(rule ccontr)auto } + { fix i :: 'a assume i:"i\Basis" + fix x::"'a" assume x:"\e>0. \x'\{x. \i\Basis. 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"]] i by auto + hence False using Basis_le_norm[OF i, of "y - x"] unfolding dist_norm inner_simps by auto } + hence "a\i \ x\i" by(rule ccontr)auto } thus ?thesis unfolding closed_limpt unfolding islimpt_approachable by blast qed -instance ordered_euclidean_space \ countable_basis_space +lemma open_box: "open (box a b)" +proof - + have "open (\i\Basis. (op \ i) -` {a \ i <..< b \ i})" + by (auto intro!: continuous_open_vimage continuous_inner continuous_at_id continuous_const) + also have "(\i\Basis. (op \ i) -` {a \ i <..< b \ i}) = box a b" + by (auto simp add: box_def inner_commute) + finally show ?thesis . +qed + +instance euclidean_space \ countable_basis_space proof - def to_cube \ "\(a, b). {Chi (real_of_rat \ op ! a)<.. op ! b)}::'a set" - def B \ "(\n. (to_cube (from_nat n)::'a set)) ` UNIV" - have "countable B" unfolding B_def by simp + def a \ "\f :: 'a \ (real \ real). \i\Basis. fst (f i) *\<^sub>R i" + then have a: "\f. (\i\Basis. fst (f i) *\<^sub>R i) = a f" by simp + def b \ "\f :: 'a \ (real \ real). \i\Basis. snd (f i) *\<^sub>R i" + then have b: "\f. (\i\Basis. snd (f i) *\<^sub>R i) = b f" by simp + def B \ "(\f. box (a f) (b f)) ` (Basis \\<^isub>E (\ \ \))" + + have "countable B" unfolding B_def + by (intro countable_image countable_PiE finite_Basis countable_SIGMA countable_rat) moreover - have "Ball B open" unfolding B_def + have "Ball B open" by (simp add: B_def open_box) + moreover have "(\A. open A \ (\B'\B. \B' = A))" proof safe - fix n show "open (to_cube (from_nat n))" - by (cases "from_nat n::rat list \ rat list") - (simp add: open_interval to_cube_def) - qed - moreover have "(\x. open x \ (\B'\B. \B' = x))" - proof safe - fix x::"'a set" assume "open x" - def lists \ "{(a, b) |a b. to_cube (a, b) \ x}" - from open_UNION[OF `open x`] - have "\(to_cube ` lists) = x" unfolding lists_def to_cube_def - by simp - moreover have "to_cube ` lists \ B" - proof - fix x assume "x \ to_cube ` lists" - then obtain l where "l \ lists" "x = to_cube l" by auto - thus "x \ B" by (auto simp add: B_def intro!: image_eqI[where x="to_nat l"]) - qed - ultimately - show "\B'\B. \B' = x" by blast + fix A::"'a set" assume "open A" + show "\B'\B. \B' = A" + apply (rule exI[of _ "{b\B. b \ A}"]) + apply (subst (3) open_UNION_box[OF `open A`]) + apply (auto simp add: a b B_def) + done qed ultimately show "\B::'a set set. countable B \ topological_basis B" unfolding topological_basis_def by blast @@ -5476,7 +5441,7 @@ text {* Intervals in general, including infinite and mixtures of open and closed. *} definition "is_interval (s::('a::euclidean_space) set) \ - (\a\s. \b\s. \x. (\i x$$i \ x$$i \ b$$i) \ (b$$i \ x$$i \ x$$i \ a$$i))) \ x \ s)" + (\a\s. \b\s. \x. (\i\Basis. ((a\i \ x\i \ x\i \ b\i) \ (b\i \ x\i \ x\i \ a\i))) \ x \ s)" lemma is_interval_interval: "is_interval {a .. b::'a::ordered_euclidean_space}" (is ?th1) "is_interval {a<..x. x $$ i)" - unfolding euclidean_component_def by (rule continuous_at_inner) - lemma closed_halfspace_le: "closed {x. inner a x \ b}" by (simp add: closed_Collect_le) @@ -5573,11 +5535,11 @@ by (simp add: closed_Collect_eq) lemma closed_halfspace_component_le: - shows "closed {x::'a::euclidean_space. x$$i \ a}" + shows "closed {x::'a::euclidean_space. x\i \ a}" by (simp add: closed_Collect_le) lemma closed_halfspace_component_ge: - shows "closed {x::'a::euclidean_space. x$$i \ a}" + shows "closed {x::'a::euclidean_space. x\i \ a}" by (simp add: closed_Collect_le) text {* Openness of halfspaces. *} @@ -5589,33 +5551,33 @@ by (simp add: open_Collect_less) lemma open_halfspace_component_lt: - shows "open {x::'a::euclidean_space. x$$i < a}" + shows "open {x::'a::euclidean_space. x\i < a}" by (simp add: open_Collect_less) lemma open_halfspace_component_gt: - shows "open {x::'a::euclidean_space. x$$i > a}" + shows "open {x::'a::euclidean_space. x\i > a}" by (simp add: open_Collect_less) text{* Instantiation for intervals on @{text ordered_euclidean_space} *} lemma eucl_lessThan_eq_halfspaces: fixes a :: "'a\ordered_euclidean_space" - shows "{..ii\Basis. {x. x \ i < a \ i})" by (auto simp: eucl_less[where 'a='a]) lemma eucl_greaterThan_eq_halfspaces: fixes a :: "'a\ordered_euclidean_space" - shows "{a<..} = (\ii\Basis. {x. a \ i < x \ i})" by (auto simp: eucl_less[where 'a='a]) lemma eucl_atMost_eq_halfspaces: fixes a :: "'a\ordered_euclidean_space" - shows "{.. a} = (\i a $$ i})" + shows "{.. a} = (\i\Basis. {x. x \ i \ a \ i})" by (auto simp: eucl_le[where 'a='a]) lemma eucl_atLeast_eq_halfspaces: fixes a :: "'a\ordered_euclidean_space" - shows "{a ..} = (\i x $$ i})" + shows "{a ..} = (\i\Basis. {x. a \ i \ x \ i})" by (auto simp: eucl_le[where 'a='a]) lemma open_eucl_lessThan[simp, intro]: @@ -5640,35 +5602,24 @@ unfolding eucl_atLeast_eq_halfspaces by (simp add: closed_INT closed_Collect_le) -lemma open_vimage_euclidean_component: "open S \ open ((\x. x $$ i) -` S)" - by (auto intro!: continuous_open_vimage) - text {* This gives a simple derivation of limit component bounds. *} lemma Lim_component_le: fixes f :: "'a \ 'b::euclidean_space" - assumes "(f ---> l) net" "\ (trivial_limit net)" "eventually (\x. f(x)$$i \ b) net" - shows "l$$i \ b" -proof- - { fix x have "x \ {x::'b. inner (basis i) x \ b} \ x$$i \ b" - unfolding euclidean_component_def by auto } note * = this - show ?thesis using Lim_in_closed_set[of "{x. inner (basis i) x \ b}" f net l] unfolding * - using closed_halfspace_le[of "(basis i)::'b" b] and assms(1,2,3) by auto -qed + assumes "(f ---> l) net" "\ (trivial_limit net)" "eventually (\x. f(x)\i \ b) net" + shows "l\i \ b" + by (rule tendsto_le[OF assms(2) tendsto_const tendsto_inner[OF assms(1) tendsto_const] assms(3)]) lemma Lim_component_ge: fixes f :: "'a \ 'b::euclidean_space" - assumes "(f ---> l) net" "\ (trivial_limit net)" "eventually (\x. b \ (f x)$$i) net" - shows "b \ l$$i" -proof- - { fix x have "x \ {x::'b. inner (basis i) x \ b} \ x$$i \ b" - unfolding euclidean_component_def by auto } note * = this - show ?thesis using Lim_in_closed_set[of "{x. inner (basis i) x \ b}" f net l] unfolding * - using closed_halfspace_ge[of b "(basis i)"] and assms(1,2,3) by auto -qed + assumes "(f ---> l) net" "\ (trivial_limit net)" "eventually (\x. b \ (f x)\i) net" + shows "b \ l\i" + by (rule tendsto_le[OF assms(2) tendsto_inner[OF assms(1) tendsto_const] tendsto_const assms(3)]) lemma Lim_component_eq: fixes f :: "'a \ 'b::euclidean_space" - assumes net:"(f ---> l) net" "~(trivial_limit net)" and ev:"eventually (\x. f(x)$$i = b) net" - shows "l$$i = b" - using ev[unfolded order_eq_iff eventually_conj_iff] using Lim_component_ge[OF net, of b i] and Lim_component_le[OF net, of i b] by auto + assumes net:"(f ---> l) net" "~(trivial_limit net)" and ev:"eventually (\x. f(x)\i = b) net" + shows "l\i = b" + using ev[unfolded order_eq_iff eventually_conj_iff] + using Lim_component_ge[OF net, of b i] and Lim_component_le[OF net, of i b] by auto + text{* Limits relative to a union. *} lemma eventually_within_Un: @@ -5726,9 +5677,8 @@ qed lemma connected_ivt_component: fixes x::"'a::euclidean_space" shows - "connected s \ x \ s \ y \ s \ x$$k \ a \ a \ y$$k \ (\z\s. z$$k = a)" - using connected_ivt_hyperplane[of s x y "(basis k)::'a" a] - unfolding euclidean_component_def by auto + "connected s \ x \ s \ y \ s \ x\k \ a \ a \ y\k \ (\z\s. z\k = a)" + using connected_ivt_hyperplane[of s x y "k::'a" a] by (auto simp: inner_commute) subsection {* Homeomorphisms *} @@ -5998,95 +5948,91 @@ subsection {* Some properties of a canonical subspace *} lemma subspace_substandard: - "subspace {x::'a::euclidean_space. (\i x$$i = 0)}" - unfolding subspace_def by auto + "subspace {x::'a::euclidean_space. (\i\Basis. P i \ x\i = 0)}" + unfolding subspace_def by (auto simp: inner_add_left) lemma closed_substandard: - "closed {x::'a::euclidean_space. \i x$$i = 0}" (is "closed ?A") + "closed {x::'a::euclidean_space. \i\Basis. P i --> x\i = 0}" (is "closed ?A") proof- - let ?D = "{i. P i} \ {..i\?D. {x::'a. x$$i = 0})" + let ?D = "{i\Basis. P i}" + have "closed (\i\?D. {x::'a. x\i = 0})" by (simp add: closed_INT closed_Collect_eq) - also have "(\i\?D. {x::'a. x$$i = 0}) = ?A" + also have "(\i\?D. {x::'a. x\i = 0}) = ?A" by auto finally show "closed ?A" . qed -lemma dim_substandard: assumes "d\{..i d \ x$$i = 0} = card d" (is "dim ?A = _") +lemma dim_substandard: assumes d: "d \ Basis" + shows "dim {x::'a::euclidean_space. \i\Basis. i \ d \ x\i = 0} = card d" (is "dim ?A = _") proof- - let ?D = "{.. ?A" by auto + let ?D = "Basis :: 'a set" + have "d \ ?A" using d by (auto simp: inner_Basis) moreover - { fix x::"'a" assume "x\?A" - hence "finite d" "x\?A" using assms by(auto intro:finite_subset) - hence "x\ span ?B" + { fix x::"'a" assume "x \ ?A" + hence "finite d" "x \ ?A" using assms by(auto intro: finite_subset[OF _ finite_Basis]) + from this d have "x \ span d" proof(induct d arbitrary: x) - case empty hence "x=0" apply(subst euclidean_eq) by auto + case empty hence "x=0" apply(rule_tac euclidean_eqI) by auto thus ?case using subspace_0[OF subspace_span[of "{}"]] by auto next case (insert k F) - hence *:"\i insert k F \ x $$ i = 0" by auto + hence *:"\i\Basis. i \ insert k F \ x \ i = 0" by auto have **:"F \ insert k F" by auto - def y \ "x - x$$k *\<^sub>R basis k" - have y:"x = y + (x$$k) *\<^sub>R basis k" unfolding y_def by auto - { fix i assume i':"i \ F" - hence "y $$ i = 0" unfolding y_def - using *[THEN spec[where x=i]] by auto } - hence "y \ span (basis ` F)" using insert(3) by auto - hence "y \ span (basis ` (insert k F))" - using span_mono[of "?bas ` F" "?bas ` (insert k F)"] - using image_mono[OF **, of basis] using assms by auto + def y \ "x - (x\k) *\<^sub>R k" + have y:"x = y + (x\k) *\<^sub>R k" unfolding y_def by auto + { fix i assume i': "i \ F" "i \ Basis" + hence "y \ i = 0" unfolding y_def + using *[THEN bspec[where x=i]] insert by (auto simp: inner_simps inner_Basis) } + hence "y \ span F" using insert by auto + hence "y \ span (insert k F)" + using span_mono[of F "insert k F"] using assms by auto moreover - have "basis k \ span (?bas ` (insert k F))" by(rule span_superset, auto) - hence "x$$k *\<^sub>R basis k \ span (?bas ` (insert k F))" + have "k \ span (insert k F)" by(rule span_superset, auto) + hence "(x\k) *\<^sub>R k \ span (insert k F)" using span_mul by auto ultimately - have "y + x$$k *\<^sub>R basis k \ span (?bas ` (insert k F))" + have "y + (x\k) *\<^sub>R k \ span (insert k F)" using span_add by auto thus ?case using y by auto qed } - hence "?A \ span ?B" by auto + hence "?A \ span d" by auto moreover - { fix x assume "x \ ?B" - hence "x\{(basis i)::'a |i. i \ ?D}" using assms by auto } - hence "independent ?B" using independent_mono[OF independent_basis, of ?B] and assms by auto + { fix x assume "x \ d" hence "x \ ?D" using assms by auto } + hence "independent d" using independent_mono[OF independent_Basis, of d] and assms by auto moreover have "d \ ?D" unfolding subset_eq using assms by auto - hence *:"inj_on (basis::nat\'a) d" using subset_inj_on[OF basis_inj, of "d"] by auto - have "card ?B = card d" unfolding card_image[OF *] by auto - ultimately show ?thesis using dim_unique[of "basis ` d" ?A] by auto + ultimately show ?thesis using dim_unique[of d ?A] by auto qed text{* Hence closure and completeness of all subspaces. *} -lemma closed_subspace_lemma: "n \ card (UNIV::'n::finite set) \ \A::'n set. card A = n" -apply (induct n) -apply (rule_tac x="{}" in exI, simp) -apply clarsimp -apply (subgoal_tac "\x. x \ A") -apply (erule exE) -apply (rule_tac x="insert x A" in exI, simp) -apply (subgoal_tac "A \ UNIV", auto) -done +lemma ex_card: assumes "n \ card A" shows "\S\A. card S = n" +proof cases + assume "finite A" + from ex_bij_betw_nat_finite[OF this] guess f .. + moreover with `n \ card A` have "{..< n} \ {..< card A}" "inj_on f {..< n}" + by (auto simp: bij_betw_def intro: subset_inj_on) + ultimately have "f ` {..< n} \ A" "card (f ` {..< n}) = n" + by (auto simp: bij_betw_def card_image) + then show ?thesis by blast +next + assume "\ finite A" with `n \ card A` show ?thesis by force +qed lemma closed_subspace: fixes s::"('a::euclidean_space) set" assumes "subspace s" shows "closed s" proof- - have *:"dim s \ DIM('a)" using dim_subset_UNIV by auto - def d \ "{..f. linear f \ f ` {x::'a. \i d \ x $$ i = 0} = s \ - inj_on f {x::'a. \i d \ x $$ i = 0}" - apply(rule subspace_isomorphism[OF subspace_substandard[of "\i. i \ d"]]) - using dim_substandard[of d,where 'a='a] and t unfolding d_def using * assms by auto - then guess f apply-by(erule exE conjE)+ note f = this + have "dim s \ card (Basis :: 'a set)" using dim_subset_UNIV by auto + with ex_card[OF this] obtain d :: "'a set" where t: "card d = dim s" and d: "d \ Basis" by auto + let ?t = "{x::'a. \i\Basis. i \ d \ x\i = 0}" + have "\f. linear f \ f ` {x::'a. \i\Basis. i \ d \ x \ i = 0} = s \ + inj_on f {x::'a. \i\Basis. i \ d \ x \ i = 0}" + using dim_substandard[of d] t d assms + by (intro subspace_isomorphism[OF subspace_substandard[of "\i. i \ d"]]) (auto simp: inner_Basis) + then guess f by (elim exE conjE) note f = this interpret f: bounded_linear f using f unfolding linear_conv_bounded_linear by auto - have "\x\?t. f x = 0 \ x = 0" using f.zero using f(3)[unfolded inj_on_def] - by(erule_tac x=0 in ballE) auto + { fix x have "x\?t \ f x = 0 \ x = 0" using f.zero d f(3)[THEN inj_onD, of x 0] by auto } moreover have "closed ?t" using closed_substandard . moreover have "subspace ?t" using subspace_substandard . ultimately show ?thesis using closed_injective_image_subspace[of ?t f] @@ -6144,7 +6090,7 @@ proof(cases "m=0") { fix x assume "x \ c" "c \ x" hence "x=c" unfolding eucl_le[where 'a='a] apply- - apply(subst euclidean_eq) by (auto intro: order_antisym) } + apply(subst euclidean_eq_iff) by (auto intro: order_antisym) } moreover case True moreover have "c \ {m *\<^sub>R a + c..m *\<^sub>R b + c}" unfolding True by(auto simp add: eucl_le[where 'a='a]) ultimately show ?thesis by auto @@ -6152,23 +6098,23 @@ case False { fix y assume "a \ y" "y \ b" "m > 0" hence "m *\<^sub>R a + c \ m *\<^sub>R y + c" "m *\<^sub>R y + c \ m *\<^sub>R b + c" - unfolding eucl_le[where 'a='a] by auto + unfolding eucl_le[where 'a='a] by (auto simp: inner_simps) } moreover { fix y assume "a \ y" "y \ b" "m < 0" hence "m *\<^sub>R b + c \ m *\<^sub>R y + c" "m *\<^sub>R y + c \ m *\<^sub>R a + c" - unfolding eucl_le[where 'a='a] by(auto simp add: mult_left_mono_neg) + unfolding eucl_le[where 'a='a] by(auto simp add: mult_left_mono_neg inner_simps) } moreover { fix y assume "m > 0" "m *\<^sub>R a + c \ y" "y \ m *\<^sub>R b + c" hence "y \ (\x. m *\<^sub>R x + c) ` {a..b}" unfolding image_iff Bex_def mem_interval eucl_le[where 'a='a] apply (intro exI[where x="(1 / m) *\<^sub>R (y - c)"]) - by(auto simp add: pos_le_divide_eq pos_divide_le_eq mult_commute diff_le_iff) + by(auto simp add: pos_le_divide_eq pos_divide_le_eq mult_commute diff_le_iff inner_simps) } moreover { fix y assume "m *\<^sub>R b + c \ y" "y \ m *\<^sub>R a + c" "m < 0" hence "y \ (\x. m *\<^sub>R x + c) ` {a..b}" unfolding image_iff Bex_def mem_interval eucl_le[where 'a='a] apply (intro exI[where x="(1 / m) *\<^sub>R (y - c)"]) - by(auto simp add: neg_le_divide_eq neg_divide_le_eq mult_commute diff_le_iff) + by(auto simp add: neg_le_divide_eq neg_divide_le_eq mult_commute diff_le_iff inner_simps) } ultimately show ?thesis using False by auto qed diff -r 46be26e02456 -r 899c9c4e4a4c src/HOL/Probability/Borel_Space.thy --- a/src/HOL/Probability/Borel_Space.thy Fri Dec 14 14:46:01 2012 +0100 +++ b/src/HOL/Probability/Borel_Space.thy Fri Dec 14 15:46:01 2012 +0100 @@ -126,344 +126,6 @@ apply(rule borel_measurableI) using continuous_open_preimage[OF assms] unfolding vimage_def by auto -section "Borel spaces on euclidean spaces" - -lemma borel_measurable_euclidean_component'[measurable]: - "(\x::'a::euclidean_space. x $$ i) \ borel_measurable borel" - by (intro continuous_on_euclidean_component continuous_on_id borel_measurable_continuous_on1) - -lemma borel_measurable_euclidean_component: - "(f :: 'a \ 'b::euclidean_space) \ borel_measurable M \(\x. f x $$ i) \ borel_measurable M" - by simp - -lemma [measurable]: - fixes a b :: "'a\ordered_euclidean_space" - shows lessThan_borel: "{..< a} \ sets borel" - and greaterThan_borel: "{a <..} \ sets borel" - and greaterThanLessThan_borel: "{a<.. sets borel" - and atMost_borel: "{..a} \ sets borel" - and atLeast_borel: "{a..} \ sets borel" - and atLeastAtMost_borel: "{a..b} \ sets borel" - and greaterThanAtMost_borel: "{a<..b} \ sets borel" - and atLeastLessThan_borel: "{a.. sets borel" - unfolding greaterThanAtMost_def atLeastLessThan_def - by (blast intro: borel_open borel_closed)+ - -lemma - shows hafspace_less_borel: "{x::'a::euclidean_space. a < x $$ i} \ sets borel" - and hafspace_greater_borel: "{x::'a::euclidean_space. x $$ i < a} \ sets borel" - and hafspace_less_eq_borel: "{x::'a::euclidean_space. a \ x $$ i} \ sets borel" - and hafspace_greater_eq_borel: "{x::'a::euclidean_space. x $$ i \ a} \ sets borel" - by simp_all - -lemma borel_measurable_less[measurable]: - fixes f :: "'a \ real" - assumes f: "f \ borel_measurable M" - assumes g: "g \ borel_measurable M" - shows "{w \ space M. f w < g w} \ sets M" -proof - - have "{w \ space M. f w < g w} = {x \ space M. \r. f x < of_rat r \ of_rat r < g x}" - using Rats_dense_in_real by (auto simp add: Rats_def) - with f g show ?thesis - by simp -qed - -lemma - fixes f :: "'a \ real" - assumes f[measurable]: "f \ borel_measurable M" - assumes g[measurable]: "g \ borel_measurable M" - shows borel_measurable_le[measurable]: "{w \ space M. f w \ g w} \ sets M" - and borel_measurable_eq[measurable]: "{w \ space M. f w = g w} \ sets M" - and borel_measurable_neq: "{w \ space M. f w \ g w} \ sets M" - unfolding eq_iff not_less[symmetric] - by measurable - -subsection "Borel space equals sigma algebras over intervals" - -lemma borel_sigma_sets_subset: - "A \ sets borel \ sigma_sets UNIV A \ sets borel" - using sets.sigma_sets_subset[of A borel] by simp - -lemma borel_eq_sigmaI1: - fixes F :: "'i \ 'a::topological_space set" and X :: "'a::topological_space set set" - assumes borel_eq: "borel = sigma UNIV X" - assumes X: "\x. x \ X \ x \ sets (sigma UNIV (range F))" - assumes F: "\i. F i \ sets borel" - shows "borel = sigma UNIV (range F)" - unfolding borel_def -proof (intro sigma_eqI antisym) - have borel_rev_eq: "sigma_sets UNIV {S::'a set. open S} = sets borel" - unfolding borel_def by simp - also have "\ = sigma_sets UNIV X" - unfolding borel_eq by simp - also have "\ \ sigma_sets UNIV (range F)" - using X by (intro sigma_algebra.sigma_sets_subset[OF sigma_algebra_sigma_sets]) auto - finally show "sigma_sets UNIV {S. open S} \ sigma_sets UNIV (range F)" . - show "sigma_sets UNIV (range F) \ sigma_sets UNIV {S. open S}" - unfolding borel_rev_eq using F by (intro borel_sigma_sets_subset) auto -qed auto - -lemma borel_eq_sigmaI2: - fixes F :: "'i \ 'j \ 'a::topological_space set" - and G :: "'l \ 'k \ 'a::topological_space set" - assumes borel_eq: "borel = sigma UNIV (range (\(i, j). G i j))" - assumes X: "\i j. G i j \ sets (sigma UNIV (range (\(i, j). F i j)))" - assumes F: "\i j. F i j \ sets borel" - shows "borel = sigma UNIV (range (\(i, j). F i j))" - using assms by (intro borel_eq_sigmaI1[where X="range (\(i, j). G i j)" and F="(\(i, j). F i j)"]) auto - -lemma borel_eq_sigmaI3: - fixes F :: "'i \ 'j \ 'a::topological_space set" and X :: "'a::topological_space set set" - assumes borel_eq: "borel = sigma UNIV X" - assumes X: "\x. x \ X \ x \ sets (sigma UNIV (range (\(i, j). F i j)))" - assumes F: "\i j. F i j \ sets borel" - shows "borel = sigma UNIV (range (\(i, j). F i j))" - using assms by (intro borel_eq_sigmaI1[where X=X and F="(\(i, j). F i j)"]) auto - -lemma borel_eq_sigmaI4: - fixes F :: "'i \ 'a::topological_space set" - and G :: "'l \ 'k \ 'a::topological_space set" - assumes borel_eq: "borel = sigma UNIV (range (\(i, j). G i j))" - assumes X: "\i j. G i j \ sets (sigma UNIV (range F))" - assumes F: "\i. F i \ sets borel" - shows "borel = sigma UNIV (range F)" - using assms by (intro borel_eq_sigmaI1[where X="range (\(i, j). G i j)" and F=F]) auto - -lemma borel_eq_sigmaI5: - fixes F :: "'i \ 'j \ 'a::topological_space set" and G :: "'l \ 'a::topological_space set" - assumes borel_eq: "borel = sigma UNIV (range G)" - assumes X: "\i. G i \ sets (sigma UNIV (range (\(i, j). F i j)))" - assumes F: "\i j. F i j \ sets borel" - shows "borel = sigma UNIV (range (\(i, j). F i j))" - using assms by (intro borel_eq_sigmaI1[where X="range G" and F="(\(i, j). F i j)"]) auto - -lemma halfspace_gt_in_halfspace: - "{x\'a. a < x $$ i} \ sigma_sets UNIV (range (\ (a, i). {x\'a\ordered_euclidean_space. x $$ i < a}))" - (is "?set \ ?SIGMA") -proof - - interpret sigma_algebra UNIV ?SIGMA - by (intro sigma_algebra_sigma_sets) simp_all - have *: "?set = (\n. UNIV - {x\'a. x $$ i < a + 1 / real (Suc n)})" - proof (safe, simp_all add: not_less) - fix x :: 'a assume "a < x $$ i" - with reals_Archimedean[of "x $$ i - a"] - obtain n where "a + 1 / real (Suc n) < x $$ i" - by (auto simp: inverse_eq_divide field_simps) - then show "\n. a + 1 / real (Suc n) \ x $$ i" - by (blast intro: less_imp_le) - next - fix x n - have "a < a + 1 / real (Suc n)" by auto - also assume "\ \ x" - finally show "a < x" . - qed - show "?set \ ?SIGMA" unfolding * - by (auto del: Diff intro!: Diff) -qed - -lemma borel_eq_halfspace_less: - "borel = sigma UNIV (range (\(a, i). {x::'a::ordered_euclidean_space. x $$ i < a}))" - (is "_ = ?SIGMA") -proof (rule borel_eq_sigmaI3[OF borel_def]) - fix S :: "'a set" assume "S \ {S. open S}" - then have "open S" by simp - from open_UNION[OF this] - obtain I where *: "S = - (\(a, b)\I. - (\ i op ! a)::'a) $$ i < x $$ i}) \ - (\ i op ! b)::'a) $$ i}))" - unfolding greaterThanLessThan_def - unfolding eucl_greaterThan_eq_halfspaces[where 'a='a] - unfolding eucl_lessThan_eq_halfspaces[where 'a='a] - by blast - show "S \ ?SIGMA" - unfolding * - by (safe intro!: sets.countable_UN sets.Int sets.countable_INT) - (auto intro!: halfspace_gt_in_halfspace) -qed auto - -lemma borel_eq_halfspace_le: - "borel = sigma UNIV (range (\ (a, i). {x::'a::ordered_euclidean_space. x $$ i \ a}))" - (is "_ = ?SIGMA") -proof (rule borel_eq_sigmaI2[OF borel_eq_halfspace_less]) - fix a i - have *: "{x::'a. x$$i < a} = (\n. {x. x$$i \ a - 1/real (Suc n)})" - proof (safe, simp_all) - fix x::'a assume *: "x$$i < a" - with reals_Archimedean[of "a - x$$i"] - obtain n where "x $$ i < a - 1 / (real (Suc n))" - by (auto simp: field_simps inverse_eq_divide) - then show "\n. x $$ i \ a - 1 / (real (Suc n))" - by (blast intro: less_imp_le) - next - fix x::'a and n - assume "x$$i \ a - 1 / real (Suc n)" - also have "\ < a" by auto - finally show "x$$i < a" . - qed - show "{x. x$$i < a} \ ?SIGMA" unfolding * - by (safe intro!: sets.countable_UN) auto -qed auto - -lemma borel_eq_halfspace_ge: - "borel = sigma UNIV (range (\ (a, i). {x\'a\ordered_euclidean_space. a \ x $$ i}))" - (is "_ = ?SIGMA") -proof (rule borel_eq_sigmaI2[OF borel_eq_halfspace_less]) - fix a i have *: "{x::'a. x$$i < a} = space ?SIGMA - {x::'a. a \ x$$i}" by auto - show "{x. x$$i < a} \ ?SIGMA" unfolding * - by (safe intro!: sets.compl_sets) auto -qed auto - -lemma borel_eq_halfspace_greater: - "borel = sigma UNIV (range (\ (a, i). {x\'a\ordered_euclidean_space. a < x $$ i}))" - (is "_ = ?SIGMA") -proof (rule borel_eq_sigmaI2[OF borel_eq_halfspace_le]) - fix a i have *: "{x::'a. x$$i \ a} = space ?SIGMA - {x::'a. a < x$$i}" by auto - show "{x. x$$i \ a} \ ?SIGMA" unfolding * - by (safe intro!: sets.compl_sets) auto -qed auto - -lemma borel_eq_atMost: - "borel = sigma UNIV (range (\a. {..a\'a\ordered_euclidean_space}))" - (is "_ = ?SIGMA") -proof (rule borel_eq_sigmaI4[OF borel_eq_halfspace_le]) - fix a i show "{x. x$$i \ a} \ ?SIGMA" - proof cases - assume "i < DIM('a)" - then have *: "{x::'a. x$$i \ a} = (\k::nat. {.. (\\ n. if n = i then a else real k)})" - proof (safe, simp_all add: eucl_le[where 'a='a] split: split_if_asm) - fix x - from real_arch_simple[of "Max ((\i. x$$i)`{..i. i < DIM('a) \ x$$i \ real k" - by (subst (asm) Max_le_iff) auto - then show "\k::nat. \ia. ia \ i \ ia < DIM('a) \ x $$ ia \ real k" - by (auto intro!: exI[of _ k]) - qed - show "{x. x$$i \ a} \ ?SIGMA" unfolding * - by (safe intro!: sets.countable_UN) auto - qed (auto intro: sigma_sets_top sigma_sets.Empty) -qed auto - -lemma borel_eq_greaterThan: - "borel = sigma UNIV (range (\a\'a\ordered_euclidean_space. {a<..}))" - (is "_ = ?SIGMA") -proof (rule borel_eq_sigmaI4[OF borel_eq_halfspace_le]) - fix a i show "{x. x$$i \ a} \ ?SIGMA" - proof cases - assume "i < DIM('a)" - have "{x::'a. x$$i \ a} = UNIV - {x::'a. a < x$$i}" by auto - also have *: "{x::'a. a < x$$i} = (\k::nat. {(\\ n. if n = i then a else -real k) <..})" using `i k::nat. \ia. ia \ i \ ia < DIM('a) \ -real k < x $$ ia" - by (auto intro!: exI[of _ k]) - qed - finally show "{x. x$$i \ a} \ ?SIGMA" - apply (simp only:) - apply (safe intro!: sets.countable_UN sets.Diff) - apply (auto intro: sigma_sets_top) - done - qed (auto intro: sigma_sets_top sigma_sets.Empty) -qed auto - -lemma borel_eq_lessThan: - "borel = sigma UNIV (range (\a\'a\ordered_euclidean_space. {.. x$$i} \ ?SIGMA" - proof cases - fix a i assume "i < DIM('a)" - have "{x::'a. a \ x$$i} = UNIV - {x::'a. x$$i < a}" by auto - also have *: "{x::'a. x$$i < a} = (\k::nat. {..< (\\ n. if n = i then a else real k)})" using `i k::nat. \ia. ia \ i \ ia < DIM('a) \ x $$ ia < real k" - by (auto intro!: exI[of _ k]) - qed - finally show "{x. a \ x$$i} \ ?SIGMA" - apply (simp only:) - apply (safe intro!: sets.countable_UN sets.Diff) - apply (auto intro: sigma_sets_top) - done - qed (auto intro: sigma_sets_top sigma_sets.Empty) -qed auto - -lemma borel_eq_atLeastAtMost: - "borel = sigma UNIV (range (\(a,b). {a..b} \'a\ordered_euclidean_space set))" - (is "_ = ?SIGMA") -proof (rule borel_eq_sigmaI5[OF borel_eq_atMost]) - fix a::'a - have *: "{..a} = (\n::nat. {- real n *\<^sub>R One .. a})" - proof (safe, simp_all add: eucl_le[where 'a='a]) - fix x - from real_arch_simple[of "Max ((\i. - x$$i)`{.. real k" - by (subst (asm) Max_le_iff) (auto simp: field_simps) - then have "- real k \ x$$i" by simp } - then show "\n::nat. \i x $$ i" - by (auto intro!: exI[of _ k]) - qed - show "{..a} \ ?SIGMA" unfolding * - by (safe intro!: sets.countable_UN) - (auto intro!: sigma_sets_top) -qed auto - -lemma borel_eq_greaterThanLessThan: - "borel = sigma UNIV (range (\ (a, b). {a <..< b} :: 'a \ ordered_euclidean_space set))" - (is "_ = ?SIGMA") -proof (rule borel_eq_sigmaI1[OF borel_def]) - fix M :: "'a set" assume "M \ {S. open S}" - then have "open M" by simp - show "M \ ?SIGMA" - apply (subst open_UNION[OF `open M`]) - apply (safe intro!: sets.countable_UN) - apply auto - done -qed auto - -lemma borel_eq_atLeastLessThan: - "borel = sigma UNIV (range (\(a, b). {a ..< b :: real}))" (is "_ = ?SIGMA") -proof (rule borel_eq_sigmaI5[OF borel_eq_lessThan]) - have move_uminus: "\x y::real. -x \ y \ -y \ x" by auto - fix x :: real - have "{..i::nat. {-real i ..< x})" - by (auto simp: move_uminus real_arch_simple) - then show "{..< x} \ ?SIGMA" - by (auto intro: sigma_sets.intros) -qed auto - -lemma borel_eq_closed: "borel = sigma UNIV (Collect closed)" - unfolding borel_def -proof (intro sigma_eqI sigma_sets_eqI, safe) - fix x :: "'a set" assume "open x" - hence "x = UNIV - (UNIV - x)" by auto - also have "\ \ sigma_sets UNIV (Collect closed)" - by (rule sigma_sets.Compl) - (auto intro!: sigma_sets.Basic simp: `open x`) - finally show "x \ sigma_sets UNIV (Collect closed)" by simp -next - fix x :: "'a set" assume "closed x" - hence "x = UNIV - (UNIV - x)" by auto - also have "\ \ sigma_sets UNIV (Collect open)" - by (rule sigma_sets.Compl) - (auto intro!: sigma_sets.Basic simp: `closed x`) - finally show "x \ sigma_sets UNIV (Collect open)" by simp -qed simp_all - lemma borel_eq_countable_basis: fixes B::"'a::topological_space set set" assumes "countable B" @@ -491,82 +153,62 @@ "borel = sigma UNIV union_closed_basis" by (rule borel_eq_countable_basis[OF countable_union_closed_basis basis_union_closed_basis]) -lemma borel_measurable_halfspacesI: - fixes f :: "'a \ 'c\ordered_euclidean_space" - assumes F: "borel = sigma UNIV (range F)" - and S_eq: "\a i. S a i = f -` F (a,i) \ space M" - and S: "\a i. \ i < DIM('c) \ S a i \ sets M" - shows "f \ borel_measurable M = (\ia::real. S a i \ sets M)" -proof safe - fix a :: real and i assume i: "i < DIM('c)" and f: "f \ borel_measurable M" - then show "S a i \ sets M" unfolding assms - by (auto intro!: measurable_sets sigma_sets.Basic simp: assms(1)) -next - assume a: "\ia. S a i \ sets M" - { fix a i have "S a i \ sets M" - proof cases - assume "i < DIM('c)" - with a show ?thesis unfolding assms(2) by simp - next - assume "\ i < DIM('c)" - from S[OF this] show ?thesis . - qed } - then show "f \ borel_measurable M" - by (auto intro!: measurable_measure_of simp: S_eq F) +lemma topological_basis_prod: + assumes A: "topological_basis A" and B: "topological_basis B" + shows "topological_basis ((\(a, b). a \ b) ` (A \ B))" + unfolding topological_basis_def +proof (safe, simp_all del: ex_simps add: subset_image_iff ex_simps(1)[symmetric]) + fix S :: "('a \ 'b) set" assume "open S" + then show "\X\A \ B. (\(a,b)\X. a \ b) = S" + proof (safe intro!: exI[of _ "{x\A \ B. fst x \ snd x \ S}"]) + fix x y assume "(x, y) \ S" + from open_prod_elim[OF `open S` this] + obtain a b where a: "open a""x \ a" and b: "open b" "y \ b" and "a \ b \ S" + by (metis mem_Sigma_iff) + moreover from topological_basisE[OF A a] guess A0 . + moreover from topological_basisE[OF B b] guess B0 . + ultimately show "(x, y) \ (\(a, b)\{X \ A \ B. fst X \ snd X \ S}. a \ b)" + by (intro UN_I[of "(A0, B0)"]) auto + qed auto +qed (metis A B topological_basis_open open_Times) + +instance prod :: (countable_basis_space, countable_basis_space) countable_basis_space +proof + obtain A :: "'a set set" where "countable A" "topological_basis A" + using ex_countable_basis by auto + moreover + obtain B :: "'b set set" where "countable B" "topological_basis B" + using ex_countable_basis by auto + ultimately show "\B::('a \ 'b) set set. countable B \ topological_basis B" + by (auto intro!: exI[of _ "(\(a, b). a \ b) ` (A \ B)"] topological_basis_prod) qed -lemma borel_measurable_iff_halfspace_le: - fixes f :: "'a \ 'c\ordered_euclidean_space" - shows "f \ borel_measurable M = (\ia. {w \ space M. f w $$ i \ a} \ sets M)" - by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_le]) auto - -lemma borel_measurable_iff_halfspace_less: - fixes f :: "'a \ 'c\ordered_euclidean_space" - shows "f \ borel_measurable M \ (\ia. {w \ space M. f w $$ i < a} \ sets M)" - by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_less]) auto - -lemma borel_measurable_iff_halfspace_ge: - fixes f :: "'a \ 'c\ordered_euclidean_space" - shows "f \ borel_measurable M = (\ia. {w \ space M. a \ f w $$ i} \ sets M)" - by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_ge]) auto - -lemma borel_measurable_iff_halfspace_greater: - fixes f :: "'a \ 'c\ordered_euclidean_space" - shows "f \ borel_measurable M \ (\ia. {w \ space M. a < f w $$ i} \ sets M)" - by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_greater]) auto - -lemma borel_measurable_iff_le: - "(f::'a \ real) \ borel_measurable M = (\a. {w \ space M. f w \ a} \ sets M)" - using borel_measurable_iff_halfspace_le[where 'c=real] by simp +lemma borel_measurable_Pair[measurable (raw)]: + fixes f :: "'a \ 'b::countable_basis_space" and g :: "'a \ 'c::countable_basis_space" + assumes f[measurable]: "f \ borel_measurable M" + assumes g[measurable]: "g \ borel_measurable M" + shows "(\x. (f x, g x)) \ borel_measurable M" +proof (subst borel_eq_countable_basis) + let ?B = "SOME B::'b set set. countable B \ topological_basis B" + let ?C = "SOME B::'c set set. countable B \ topological_basis B" + let ?P = "(\(b, c). b \ c) ` (?B \ ?C)" + show "countable ?P" "topological_basis ?P" + by (auto intro!: countable_basis topological_basis_prod is_basis) -lemma borel_measurable_iff_less: - "(f::'a \ real) \ borel_measurable M = (\a. {w \ space M. f w < a} \ sets M)" - using borel_measurable_iff_halfspace_less[where 'c=real] by simp - -lemma borel_measurable_iff_ge: - "(f::'a \ real) \ borel_measurable M = (\a. {w \ space M. a \ f w} \ sets M)" - using borel_measurable_iff_halfspace_ge[where 'c=real] - by simp - -lemma borel_measurable_iff_greater: - "(f::'a \ real) \ borel_measurable M = (\a. {w \ space M. a < f w} \ sets M)" - using borel_measurable_iff_halfspace_greater[where 'c=real] by simp - -lemma borel_measurable_euclidean_space: - fixes f :: "'a \ 'c::ordered_euclidean_space" - shows "f \ borel_measurable M \ (\ix. f x $$ i) \ borel_measurable M)" -proof safe - fix i assume "f \ borel_measurable M" - then show "(\x. f x $$ i) \ borel_measurable M" - by (auto intro: borel_measurable_euclidean_component) -next - assume f: "\ix. f x $$ i) \ borel_measurable M" - then show "f \ borel_measurable M" - unfolding borel_measurable_iff_halfspace_le by auto + show "(\x. (f x, g x)) \ measurable M (sigma UNIV ?P)" + proof (rule measurable_measure_of) + fix S assume "S \ ?P" + then obtain b c where "b \ ?B" "c \ ?C" and S: "S = b \ c" by auto + then have borel: "open b" "open c" + by (auto intro: is_basis topological_basis_open) + have "(\x. (f x, g x)) -` S \ space M = (f -` b \ space M) \ (g -` c \ space M)" + unfolding S by auto + also have "\ \ sets M" + using borel by simp + finally show "(\x. (f x, g x)) -` S \ space M \ sets M" . + qed auto qed -subsection "Borel measurable operators" - lemma borel_measurable_continuous_on: fixes f :: "'a::topological_space \ 'b::topological_space" assumes f: "continuous_on UNIV f" and g: "g \ borel_measurable M" @@ -598,30 +240,6 @@ using measurable_comp[OF g borel_measurable_continuous_on_open'[OF cont], of c] by (simp add: comp_def) -lemma borel_measurable_uminus[measurable (raw)]: - fixes g :: "'a \ real" - assumes g: "g \ borel_measurable M" - shows "(\x. - g x) \ borel_measurable M" - by (rule borel_measurable_continuous_on[OF _ g]) (auto intro: continuous_on_minus continuous_on_id) - -lemma euclidean_component_prod: - fixes x :: "'a :: euclidean_space \ 'b :: euclidean_space" - shows "x $$ i = (if i < DIM('a) then fst x $$ i else snd x $$ (i - DIM('a)))" - unfolding euclidean_component_def basis_prod_def inner_prod_def by auto - -lemma borel_measurable_Pair[measurable (raw)]: - fixes f :: "'a \ 'b::ordered_euclidean_space" and g :: "'a \ 'c::ordered_euclidean_space" - assumes f: "f \ borel_measurable M" - assumes g: "g \ borel_measurable M" - shows "(\x. (f x, g x)) \ borel_measurable M" -proof (intro borel_measurable_iff_halfspace_le[THEN iffD2] allI impI) - fix i and a :: real assume i: "i < DIM('b \ 'c)" - have [simp]: "\P A B C. {w. (P \ A w \ B w) \ (\ P \ A w \ C w)} = - {w. A w \ (P \ B w) \ (\ P \ C w)}" by auto - from i f g show "{w \ space M. (f w, g w) $$ i \ a} \ sets M" - by (auto simp: euclidean_component_prod) -qed - lemma continuous_on_fst: "continuous_on UNIV fst" proof - have [simp]: "range fst = UNIV" by (auto simp: image_iff) @@ -639,7 +257,7 @@ qed lemma borel_measurable_continuous_Pair: - fixes f :: "'a \ 'b::ordered_euclidean_space" and g :: "'a \ 'c::ordered_euclidean_space" + fixes f :: "'a \ 'b::countable_basis_space" and g :: "'a \ 'c::countable_basis_space" assumes [measurable]: "f \ borel_measurable M" assumes [measurable]: "g \ borel_measurable M" assumes H: "continuous_on UNIV (\x. H (fst x) (snd x))" @@ -650,6 +268,413 @@ unfolding eq by (rule borel_measurable_continuous_on[OF H]) auto qed +section "Borel spaces on euclidean spaces" + +lemma borel_measurable_inner[measurable (raw)]: + fixes f g :: "'a \ 'b::{countable_basis_space, real_inner}" + assumes "f \ borel_measurable M" + assumes "g \ borel_measurable M" + shows "(\x. f x \ g x) \ borel_measurable M" + using assms + by (rule borel_measurable_continuous_Pair) + (intro continuous_on_inner continuous_on_snd continuous_on_fst) + +lemma [measurable]: + fixes a b :: "'a\ordered_euclidean_space" + shows lessThan_borel: "{..< a} \ sets borel" + and greaterThan_borel: "{a <..} \ sets borel" + and greaterThanLessThan_borel: "{a<.. sets borel" + and atMost_borel: "{..a} \ sets borel" + and atLeast_borel: "{a..} \ sets borel" + and atLeastAtMost_borel: "{a..b} \ sets borel" + and greaterThanAtMost_borel: "{a<..b} \ sets borel" + and atLeastLessThan_borel: "{a.. sets borel" + unfolding greaterThanAtMost_def atLeastLessThan_def + by (blast intro: borel_open borel_closed)+ + +lemma borel_measurable_less[measurable]: + fixes f :: "'a \ real" + assumes f: "f \ borel_measurable M" + assumes g: "g \ borel_measurable M" + shows "{w \ space M. f w < g w} \ sets M" +proof - + have "{w \ space M. f w < g w} = {x \ space M. \r. f x < of_rat r \ of_rat r < g x}" + using Rats_dense_in_real by (auto simp add: Rats_def) + with f g show ?thesis + by simp +qed + +lemma + fixes f :: "'a \ real" + assumes f[measurable]: "f \ borel_measurable M" + assumes g[measurable]: "g \ borel_measurable M" + shows borel_measurable_le[measurable]: "{w \ space M. f w \ g w} \ sets M" + and borel_measurable_eq[measurable]: "{w \ space M. f w = g w} \ sets M" + and borel_measurable_neq: "{w \ space M. f w \ g w} \ sets M" + unfolding eq_iff not_less[symmetric] + by measurable + +lemma + shows hafspace_less_borel: "{x::'a::euclidean_space. a < x \ i} \ sets borel" + and hafspace_greater_borel: "{x::'a::euclidean_space. x \ i < a} \ sets borel" + and hafspace_less_eq_borel: "{x::'a::euclidean_space. a \ x \ i} \ sets borel" + and hafspace_greater_eq_borel: "{x::'a::euclidean_space. x \ i \ a} \ sets borel" + by simp_all + +subsection "Borel space equals sigma algebras over intervals" + +lemma borel_sigma_sets_subset: + "A \ sets borel \ sigma_sets UNIV A \ sets borel" + using sets.sigma_sets_subset[of A borel] by simp + +lemma borel_eq_sigmaI1: + fixes F :: "'i \ 'a::topological_space set" and X :: "'a::topological_space set set" + assumes borel_eq: "borel = sigma UNIV X" + assumes X: "\x. x \ X \ x \ sets (sigma UNIV (F ` A))" + assumes F: "\i. i \ A \ F i \ sets borel" + shows "borel = sigma UNIV (F ` A)" + unfolding borel_def +proof (intro sigma_eqI antisym) + have borel_rev_eq: "sigma_sets UNIV {S::'a set. open S} = sets borel" + unfolding borel_def by simp + also have "\ = sigma_sets UNIV X" + unfolding borel_eq by simp + also have "\ \ sigma_sets UNIV (F`A)" + using X by (intro sigma_algebra.sigma_sets_subset[OF sigma_algebra_sigma_sets]) auto + finally show "sigma_sets UNIV {S. open S} \ sigma_sets UNIV (F`A)" . + show "sigma_sets UNIV (F`A) \ sigma_sets UNIV {S. open S}" + unfolding borel_rev_eq using F by (intro borel_sigma_sets_subset) auto +qed auto + +lemma borel_eq_sigmaI2: + fixes F :: "'i \ 'j \ 'a::topological_space set" + and G :: "'l \ 'k \ 'a::topological_space set" + assumes borel_eq: "borel = sigma UNIV ((\(i, j). G i j)`B)" + assumes X: "\i j. (i, j) \ B \ G i j \ sets (sigma UNIV ((\(i, j). F i j) ` A))" + assumes F: "\i j. (i, j) \ A \ F i j \ sets borel" + shows "borel = sigma UNIV ((\(i, j). F i j) ` A)" + using assms + by (intro borel_eq_sigmaI1[where X="(\(i, j). G i j) ` B" and F="(\(i, j). F i j)"]) auto + +lemma borel_eq_sigmaI3: + fixes F :: "'i \ 'j \ 'a::topological_space set" and X :: "'a::topological_space set set" + assumes borel_eq: "borel = sigma UNIV X" + assumes X: "\x. x \ X \ x \ sets (sigma UNIV ((\(i, j). F i j) ` A))" + assumes F: "\i j. (i, j) \ A \ F i j \ sets borel" + shows "borel = sigma UNIV ((\(i, j). F i j) ` A)" + using assms by (intro borel_eq_sigmaI1[where X=X and F="(\(i, j). F i j)"]) auto + +lemma borel_eq_sigmaI4: + fixes F :: "'i \ 'a::topological_space set" + and G :: "'l \ 'k \ 'a::topological_space set" + assumes borel_eq: "borel = sigma UNIV ((\(i, j). G i j)`A)" + assumes X: "\i j. (i, j) \ A \ G i j \ sets (sigma UNIV (range F))" + assumes F: "\i. F i \ sets borel" + shows "borel = sigma UNIV (range F)" + using assms by (intro borel_eq_sigmaI1[where X="(\(i, j). G i j) ` A" and F=F]) auto + +lemma borel_eq_sigmaI5: + fixes F :: "'i \ 'j \ 'a::topological_space set" and G :: "'l \ 'a::topological_space set" + assumes borel_eq: "borel = sigma UNIV (range G)" + assumes X: "\i. G i \ sets (sigma UNIV (range (\(i, j). F i j)))" + assumes F: "\i j. F i j \ sets borel" + shows "borel = sigma UNIV (range (\(i, j). F i j))" + using assms by (intro borel_eq_sigmaI1[where X="range G" and F="(\(i, j). F i j)"]) auto + +lemma borel_eq_box: + "borel = sigma UNIV (range (\ (a, b). box a b :: 'a \ euclidean_space set))" + (is "_ = ?SIGMA") +proof (rule borel_eq_sigmaI1[OF borel_def]) + fix M :: "'a set" assume "M \ {S. open S}" + then have "open M" by simp + show "M \ ?SIGMA" + apply (subst open_UNION_box[OF `open M`]) + apply (safe intro!: sets.countable_UN' countable_PiE countable_Collect) + apply (auto intro: countable_rat) + done +qed (auto simp: box_def) + +lemma borel_eq_greaterThanLessThan: + "borel = sigma UNIV (range (\ (a, b). {a <..< b} :: 'a \ ordered_euclidean_space set))" + unfolding borel_eq_box apply (rule arg_cong2[where f=sigma]) + by (auto simp: box_def image_iff mem_interval set_eq_iff simp del: greaterThanLessThan_iff) + +lemma halfspace_gt_in_halfspace: + assumes i: "i \ A" + shows "{x\'a. a < x \ i} \ + sigma_sets UNIV ((\ (a, i). {x\'a\euclidean_space. x \ i < a}) ` (UNIV \ A))" + (is "?set \ ?SIGMA") +proof - + interpret sigma_algebra UNIV ?SIGMA + by (intro sigma_algebra_sigma_sets) simp_all + have *: "?set = (\n. UNIV - {x\'a. x \ i < a + 1 / real (Suc n)})" + proof (safe, simp_all add: not_less) + fix x :: 'a assume "a < x \ i" + with reals_Archimedean[of "x \ i - a"] + obtain n where "a + 1 / real (Suc n) < x \ i" + by (auto simp: inverse_eq_divide field_simps) + then show "\n. a + 1 / real (Suc n) \ x \ i" + by (blast intro: less_imp_le) + next + fix x n + have "a < a + 1 / real (Suc n)" by auto + also assume "\ \ x" + finally show "a < x" . + qed + show "?set \ ?SIGMA" unfolding * + by (auto del: Diff intro!: Diff i) +qed + +lemma borel_eq_halfspace_less: + "borel = sigma UNIV ((\(a, i). {x::'a::euclidean_space. x \ i < a}) ` (UNIV \ Basis))" + (is "_ = ?SIGMA") +proof (rule borel_eq_sigmaI2[OF borel_eq_box]) + fix a b :: 'a + have "box a b = {x\space ?SIGMA. \i\Basis. a \ i < x \ i \ x \ i < b \ i}" + by (auto simp: box_def) + also have "\ \ sets ?SIGMA" + by (intro sets.sets_Collect_conj sets.sets_Collect_finite_All sets.sets_Collect_const) + (auto intro!: halfspace_gt_in_halfspace countable_PiE countable_rat) + finally show "box a b \ sets ?SIGMA" . +qed auto + +lemma borel_eq_halfspace_le: + "borel = sigma UNIV ((\ (a, i). {x::'a::euclidean_space. x \ i \ a}) ` (UNIV \ Basis))" + (is "_ = ?SIGMA") +proof (rule borel_eq_sigmaI2[OF borel_eq_halfspace_less]) + fix a :: real and i :: 'a assume "(a, i) \ UNIV \ Basis" + then have i: "i \ Basis" by auto + have *: "{x::'a. x\i < a} = (\n. {x. x\i \ a - 1/real (Suc n)})" + proof (safe, simp_all) + fix x::'a assume *: "x\i < a" + with reals_Archimedean[of "a - x\i"] + obtain n where "x \ i < a - 1 / (real (Suc n))" + by (auto simp: field_simps inverse_eq_divide) + then show "\n. x \ i \ a - 1 / (real (Suc n))" + by (blast intro: less_imp_le) + next + fix x::'a and n + assume "x\i \ a - 1 / real (Suc n)" + also have "\ < a" by auto + finally show "x\i < a" . + qed + show "{x. x\i < a} \ ?SIGMA" unfolding * + by (safe intro!: sets.countable_UN) (auto intro: i) +qed auto + +lemma borel_eq_halfspace_ge: + "borel = sigma UNIV ((\ (a, i). {x\'a\euclidean_space. a \ x \ i}) ` (UNIV \ Basis))" + (is "_ = ?SIGMA") +proof (rule borel_eq_sigmaI2[OF borel_eq_halfspace_less]) + fix a :: real and i :: 'a assume i: "(a, i) \ UNIV \ Basis" + have *: "{x::'a. x\i < a} = space ?SIGMA - {x::'a. a \ x\i}" by auto + show "{x. x\i < a} \ ?SIGMA" unfolding * + using i by (safe intro!: sets.compl_sets) auto +qed auto + +lemma borel_eq_halfspace_greater: + "borel = sigma UNIV ((\ (a, i). {x\'a\euclidean_space. a < x \ i}) ` (UNIV \ Basis))" + (is "_ = ?SIGMA") +proof (rule borel_eq_sigmaI2[OF borel_eq_halfspace_le]) + fix a :: real and i :: 'a assume "(a, i) \ (UNIV \ Basis)" + then have i: "i \ Basis" by auto + have *: "{x::'a. x\i \ a} = space ?SIGMA - {x::'a. a < x\i}" by auto + show "{x. x\i \ a} \ ?SIGMA" unfolding * + by (safe intro!: sets.compl_sets) (auto intro: i) +qed auto + +lemma borel_eq_atMost: + "borel = sigma UNIV (range (\a. {..a\'a\ordered_euclidean_space}))" + (is "_ = ?SIGMA") +proof (rule borel_eq_sigmaI4[OF borel_eq_halfspace_le]) + fix a :: real and i :: 'a assume "(a, i) \ UNIV \ Basis" + then have "i \ Basis" by auto + then have *: "{x::'a. x\i \ a} = (\k::nat. {.. (\n\Basis. (if n = i then a else real k)*\<^sub>R n)})" + proof (safe, simp_all add: eucl_le[where 'a='a] split: split_if_asm) + fix x :: 'a + from real_arch_simple[of "Max ((\i. x\i)`Basis)"] guess k::nat .. + then have "\i. i \ Basis \ x\i \ real k" + by (subst (asm) Max_le_iff) auto + then show "\k::nat. \ia\Basis. ia \ i \ x \ ia \ real k" + by (auto intro!: exI[of _ k]) + qed + show "{x. x\i \ a} \ ?SIGMA" unfolding * + by (safe intro!: sets.countable_UN) auto +qed auto + +lemma borel_eq_greaterThan: + "borel = sigma UNIV (range (\a\'a\ordered_euclidean_space. {a<..}))" + (is "_ = ?SIGMA") +proof (rule borel_eq_sigmaI4[OF borel_eq_halfspace_le]) + fix a :: real and i :: 'a assume "(a, i) \ UNIV \ Basis" + then have i: "i \ Basis" by auto + have "{x::'a. x\i \ a} = UNIV - {x::'a. a < x\i}" by auto + also have *: "{x::'a. a < x\i} = + (\k::nat. {\n\Basis. (if n = i then a else -real k) *\<^sub>R n <..})" using i + proof (safe, simp_all add: eucl_less[where 'a='a] split: split_if_asm) + fix x :: 'a + from reals_Archimedean2[of "Max ((\i. -x\i)`Basis)"] + guess k::nat .. note k = this + { fix i :: 'a assume "i \ Basis" + then have "-x\i < real k" + using k by (subst (asm) Max_less_iff) auto + then have "- real k < x\i" by simp } + then show "\k::nat. \ia\Basis. ia \ i \ -real k < x \ ia" + by (auto intro!: exI[of _ k]) + qed + finally show "{x. x\i \ a} \ ?SIGMA" + apply (simp only:) + apply (safe intro!: sets.countable_UN sets.Diff) + apply (auto intro: sigma_sets_top) + done +qed auto + +lemma borel_eq_lessThan: + "borel = sigma UNIV (range (\a\'a\ordered_euclidean_space. {.. UNIV \ Basis" + then have i: "i \ Basis" by auto + have "{x::'a. a \ x\i} = UNIV - {x::'a. x\i < a}" by auto + also have *: "{x::'a. x\i < a} = (\k::nat. {..< \n\Basis. (if n = i then a else real k) *\<^sub>R n})" using `i\ Basis` + proof (safe, simp_all add: eucl_less[where 'a='a] split: split_if_asm) + fix x :: 'a + from reals_Archimedean2[of "Max ((\i. x\i)`Basis)"] + guess k::nat .. note k = this + { fix i :: 'a assume "i \ Basis" + then have "x\i < real k" + using k by (subst (asm) Max_less_iff) auto + then have "x\i < real k" by simp } + then show "\k::nat. \ia\Basis. ia \ i \ x \ ia < real k" + by (auto intro!: exI[of _ k]) + qed + finally show "{x. a \ x\i} \ ?SIGMA" + apply (simp only:) + apply (safe intro!: sets.countable_UN sets.Diff) + apply (auto intro: sigma_sets_top) + done +qed auto + +lemma borel_eq_atLeastAtMost: + "borel = sigma UNIV (range (\(a,b). {a..b} \'a\ordered_euclidean_space set))" + (is "_ = ?SIGMA") +proof (rule borel_eq_sigmaI5[OF borel_eq_atMost]) + fix a::'a + have *: "{..a} = (\n::nat. {- real n *\<^sub>R One .. a})" + proof (safe, simp_all add: eucl_le[where 'a='a]) + fix x :: 'a + from real_arch_simple[of "Max ((\i. - x\i)`Basis)"] + guess k::nat .. note k = this + { fix i :: 'a assume "i \ Basis" + with k have "- x\i \ real k" + by (subst (asm) Max_le_iff) (auto simp: field_simps) + then have "- real k \ x\i" by simp } + then show "\n::nat. \i\Basis. - real n \ x \ i" + by (auto intro!: exI[of _ k]) + qed + show "{..a} \ ?SIGMA" unfolding * + by (safe intro!: sets.countable_UN) + (auto intro!: sigma_sets_top) +qed auto + +lemma borel_eq_atLeastLessThan: + "borel = sigma UNIV (range (\(a, b). {a ..< b :: real}))" (is "_ = ?SIGMA") +proof (rule borel_eq_sigmaI5[OF borel_eq_lessThan]) + have move_uminus: "\x y::real. -x \ y \ -y \ x" by auto + fix x :: real + have "{..i::nat. {-real i ..< x})" + by (auto simp: move_uminus real_arch_simple) + then show "{..< x} \ ?SIGMA" + by (auto intro: sigma_sets.intros) +qed auto + +lemma borel_eq_closed: "borel = sigma UNIV (Collect closed)" + unfolding borel_def +proof (intro sigma_eqI sigma_sets_eqI, safe) + fix x :: "'a set" assume "open x" + hence "x = UNIV - (UNIV - x)" by auto + also have "\ \ sigma_sets UNIV (Collect closed)" + by (rule sigma_sets.Compl) + (auto intro!: sigma_sets.Basic simp: `open x`) + finally show "x \ sigma_sets UNIV (Collect closed)" by simp +next + fix x :: "'a set" assume "closed x" + hence "x = UNIV - (UNIV - x)" by auto + also have "\ \ sigma_sets UNIV (Collect open)" + by (rule sigma_sets.Compl) + (auto intro!: sigma_sets.Basic simp: `closed x`) + finally show "x \ sigma_sets UNIV (Collect open)" by simp +qed simp_all + +lemma borel_measurable_halfspacesI: + fixes f :: "'a \ 'c\euclidean_space" + assumes F: "borel = sigma UNIV (F ` (UNIV \ Basis))" + and S_eq: "\a i. S a i = f -` F (a,i) \ space M" + shows "f \ borel_measurable M = (\i\Basis. \a::real. S a i \ sets M)" +proof safe + fix a :: real and i :: 'b assume i: "i \ Basis" and f: "f \ borel_measurable M" + then show "S a i \ sets M" unfolding assms + by (auto intro!: measurable_sets simp: assms(1)) +next + assume a: "\i\Basis. \a. S a i \ sets M" + then show "f \ borel_measurable M" + by (auto intro!: measurable_measure_of simp: S_eq F) +qed + +lemma borel_measurable_iff_halfspace_le: + fixes f :: "'a \ 'c\euclidean_space" + shows "f \ borel_measurable M = (\i\Basis. \a. {w \ space M. f w \ i \ a} \ sets M)" + by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_le]) auto + +lemma borel_measurable_iff_halfspace_less: + fixes f :: "'a \ 'c\euclidean_space" + shows "f \ borel_measurable M \ (\i\Basis. \a. {w \ space M. f w \ i < a} \ sets M)" + by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_less]) auto + +lemma borel_measurable_iff_halfspace_ge: + fixes f :: "'a \ 'c\euclidean_space" + shows "f \ borel_measurable M = (\i\Basis. \a. {w \ space M. a \ f w \ i} \ sets M)" + by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_ge]) auto + +lemma borel_measurable_iff_halfspace_greater: + fixes f :: "'a \ 'c\euclidean_space" + shows "f \ borel_measurable M \ (\i\Basis. \a. {w \ space M. a < f w \ i} \ sets M)" + by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_greater]) auto + +lemma borel_measurable_iff_le: + "(f::'a \ real) \ borel_measurable M = (\a. {w \ space M. f w \ a} \ sets M)" + using borel_measurable_iff_halfspace_le[where 'c=real] by simp + +lemma borel_measurable_iff_less: + "(f::'a \ real) \ borel_measurable M = (\a. {w \ space M. f w < a} \ sets M)" + using borel_measurable_iff_halfspace_less[where 'c=real] by simp + +lemma borel_measurable_iff_ge: + "(f::'a \ real) \ borel_measurable M = (\a. {w \ space M. a \ f w} \ sets M)" + using borel_measurable_iff_halfspace_ge[where 'c=real] + by simp + +lemma borel_measurable_iff_greater: + "(f::'a \ real) \ borel_measurable M = (\a. {w \ space M. a < f w} \ sets M)" + using borel_measurable_iff_halfspace_greater[where 'c=real] by simp + +lemma borel_measurable_euclidean_space: + fixes f :: "'a \ 'c::euclidean_space" + shows "f \ borel_measurable M \ (\i\Basis. (\x. f x \ i) \ borel_measurable M)" +proof safe + assume f: "\i\Basis. (\x. f x \ i) \ borel_measurable M" + then show "f \ borel_measurable M" + by (subst borel_measurable_iff_halfspace_le) auto +qed auto + +subsection "Borel measurable operators" + +lemma borel_measurable_uminus[measurable (raw)]: + fixes g :: "'a \ real" + assumes g: "g \ borel_measurable M" + shows "(\x. - g x) \ borel_measurable M" + by (rule borel_measurable_continuous_on[OF _ g]) (auto intro: continuous_on_minus continuous_on_id) + lemma borel_measurable_add[measurable (raw)]: fixes f g :: "'a \ 'c::ordered_euclidean_space" assumes f: "f \ borel_measurable M" @@ -772,7 +797,7 @@ lemma borel_measurable_nth[measurable (raw)]: "(\x::real^'n. x $ i) \ borel_measurable borel" - by (simp add: nth_conv_component) + by (simp add: cart_eq_inner_axis) lemma convex_measurable: fixes a b :: real diff -r 46be26e02456 -r 899c9c4e4a4c src/HOL/Probability/Lebesgue_Measure.thy --- a/src/HOL/Probability/Lebesgue_Measure.thy Fri Dec 14 14:46:01 2012 +0100 +++ b/src/HOL/Probability/Lebesgue_Measure.thy Fri Dec 14 15:46:01 2012 +0100 @@ -36,7 +36,7 @@ subsection {* Standard Cubes *} definition cube :: "nat \ 'a::ordered_euclidean_space set" where - "cube n \ {\\ i. - real n .. \\ i. real n}" + "cube n \ {\i\Basis. - n *\<^sub>R i .. \i\Basis. n *\<^sub>R i}" lemma borel_cube[intro]: "cube n \ sets borel" unfolding cube_def by auto @@ -45,53 +45,44 @@ unfolding cube_def by auto lemma cube_subset[intro]: "n \ N \ cube n \ cube N" - by (fastforce simp: eucl_le[where 'a='a] cube_def) + by (fastforce simp: eucl_le[where 'a='a] cube_def setsum_negf) + +lemma cube_subset_iff: "cube n \ cube N \ n \ N" + unfolding cube_def subset_interval by (simp add: setsum_negf ex_in_conv) -lemma cube_subset_iff: - "cube n \ cube N \ n \ N" -proof - assume subset: "cube n \ (cube N::'a set)" - then have "((\\ i. real n)::'a) \ cube N" - using DIM_positive[where 'a='a] - by (fastforce simp: cube_def eucl_le[where 'a='a]) - then show "n \ N" - by (fastforce simp: cube_def eucl_le[where 'a='a]) -next - assume "n \ N" then show "cube n \ (cube N::'a set)" by (rule cube_subset) -qed - -lemma ball_subset_cube:"ball (0::'a::ordered_euclidean_space) (real n) \ cube n" - unfolding cube_def subset_eq mem_interval apply safe unfolding euclidean_lambda_beta' -proof- fix x::'a and i assume as:"x\ball 0 (real n)" "i x $$ i" "real n \ x $$ i" - using component_le_norm[of x i] by(auto simp: dist_norm) +lemma ball_subset_cube: "ball (0::'a::ordered_euclidean_space) (real n) \ cube n" + apply (simp add: cube_def subset_eq mem_interval setsum_negf eucl_le[where 'a='a]) +proof safe + fix x i :: 'a assume x: "x \ ball 0 (real n)" and i: "i \ Basis" + thus "- real n \ x \ i" "real n \ x \ i" + using Basis_le_norm[OF i, of x] by(auto simp: dist_norm) qed lemma mem_big_cube: obtains n where "x \ cube n" -proof- from reals_Archimedean2[of "norm x"] guess n .. - thus ?thesis apply-apply(rule that[where n=n]) - apply(rule ball_subset_cube[unfolded subset_eq,rule_format]) - by (auto simp add:dist_norm) +proof - + from reals_Archimedean2[of "norm x"] guess n .. + with ball_subset_cube[unfolded subset_eq, of n] + show ?thesis + by (intro that[where n=n]) (auto simp add: dist_norm) qed lemma cube_subset_Suc[intro]: "cube n \ cube (Suc n)" - unfolding cube_def subset_eq apply safe unfolding mem_interval apply auto done + unfolding cube_def subset_interval by (simp add: setsum_negf) lemma has_integral_interval_cube: fixes a b :: "'a::ordered_euclidean_space" - shows "(indicator {a .. b} has_integral - content ({\\ i. max (- real n) (a $$ i) .. \\ i. min (real n) (b $$ i)} :: 'a set)) (cube n)" + shows "(indicator {a .. b} has_integral content ({a .. b} \ cube n)) (cube n)" (is "(?I has_integral content ?R) (cube n)") proof - - let "{?N .. ?P}" = ?R have [simp]: "(\x. if x \ cube n then ?I x else 0) = indicator ?R" by (auto simp: indicator_def cube_def fun_eq_iff eucl_le[where 'a='a]) have "(?I has_integral content ?R) (cube n) \ (indicator ?R has_integral content ?R) UNIV" unfolding has_integral_restrict_univ[where s="cube n", symmetric] by simp - also have "\ \ ((\x. 1) has_integral content ?R) ?R" - unfolding indicator_def [abs_def] has_integral_restrict_univ .. - finally show ?thesis - using has_integral_const[of "1::real" "?N" "?P"] by simp + also have "\ \ ((\x. 1::real) has_integral content ?R *\<^sub>R 1) ?R" + unfolding indicator_def [abs_def] has_integral_restrict_univ real_scaleR_def mult_1_right .. + also have "((\x. 1) has_integral content ?R *\<^sub>R 1) ?R" + unfolding cube_def inter_interval by (rule has_integral_const) + finally show ?thesis . qed subsection {* Lebesgue measure *} @@ -219,11 +210,8 @@ also have "\ \ sets lebesgue" proof (safe intro!: sets.sigma_sets_subset lebesgueI) fix n :: nat and a b :: 'a - let ?N = "\\ i. max (- real n) (a $$ i)" - let ?P = "\\ i. min (real n) (b $$ i)" show "(indicator {a..b} :: 'a\real) integrable_on cube n" - unfolding integrable_on_def - using has_integral_interval_cube[of a b] by auto + unfolding integrable_on_def using has_integral_interval_cube[of a b] by auto qed finally show ?thesis . qed @@ -418,12 +406,13 @@ case (Suc n') have "real n \ (2 * real n)^1" by auto also have "(2 * real n)^1 \ (2 * real n) ^ DIM('a)" - using Suc DIM_positive[where 'a='a] by (intro power_increasing) (auto simp: real_of_nat_Suc) + using Suc DIM_positive[where 'a='a] + by (intro power_increasing) (auto simp: real_of_nat_Suc simp del: DIM_positive) finally show ?thesis . qed } ultimately show "ereal (real n) \ ereal (integral (cube n) (indicator UNIV::'a\real))" using integral_const DIM_positive[where 'a='a] - by (auto simp: cube_def content_closed_interval_cases setprod_constant) + by (auto simp: cube_def content_closed_interval_cases setprod_constant setsum_negf) qed simp lemma lmeasure_complete: "A \ B \ B \ null_sets lebesgue \ A \ null_sets lebesgue" @@ -901,24 +890,24 @@ subsection {* Equivalence between product spaces and euclidean spaces *} -definition e2p :: "'a::ordered_euclidean_space \ (nat \ real)" where - "e2p x = (\i\{.. ('a \ real)" where + "e2p x = (\i\Basis. x \ i)" -definition p2e :: "(nat \ real) \ 'a::ordered_euclidean_space" where - "p2e x = (\\ i. x i)" +definition p2e :: "('a \ real) \ 'a::ordered_euclidean_space" where + "p2e x = (\i\Basis. x i *\<^sub>R i)" lemma e2p_p2e[simp]: - "x \ extensional {.. e2p (p2e x::'a::ordered_euclidean_space) = x" + "x \ extensional Basis \ e2p (p2e x::'a::ordered_euclidean_space) = x" by (auto simp: fun_eq_iff extensional_def p2e_def e2p_def) lemma p2e_e2p[simp]: "p2e (e2p x) = (x::'a::ordered_euclidean_space)" - by (auto simp: euclidean_eq[where 'a='a] p2e_def e2p_def) + by (auto simp: euclidean_eq_iff[where 'a='a] p2e_def e2p_def) interpretation lborel_product: product_sigma_finite "\x. lborel::real measure" by default -interpretation lborel_space: finite_product_sigma_finite "\x. lborel::real measure" "{..x. lborel::real measure" "Basis" by default auto lemma sets_product_borel: @@ -930,13 +919,13 @@ qed (auto simp: borel_eq_lessThan reals_Archimedean2) lemma measurable_e2p[measurable]: - "e2p \ measurable (borel::'a::ordered_euclidean_space measure) (\\<^isub>M i\{.. measurable (borel::'a::ordered_euclidean_space measure) (\\<^isub>M (i::'a)\Basis. (lborel :: real measure))" proof (rule measurable_sigma_sets[OF sets_product_borel]) - fix A :: "(nat \ real) set" assume "A \ {\\<^isub>E i\{..\<^isub>E i\{..\ i. x i) :: 'a}" + fix A :: "('a \ real) set" assume "A \ {\\<^isub>E (i::'a)\Basis. {..\<^isub>E (i::'a)\Basis. {..i\Basis. x i *\<^sub>R i) :: 'a}" using DIM_positive by (auto simp add: set_eq_iff e2p_def - euclidean_eq[where 'a='a] eucl_less[where 'a='a]) + euclidean_eq_iff[where 'a='a] eucl_less[where 'a='a]) then show "e2p -` A \ space (borel::'a measure) \ sets borel" by simp qed (auto simp: e2p_def) @@ -945,34 +934,34 @@ lemma lborelD[measurable (raw)]: "A \ sets borel \ A \ sets lborel" by simp lemma measurable_p2e[measurable]: - "p2e \ measurable (\\<^isub>M i\{.. measurable (\\<^isub>M (i::'a)\Basis. (lborel :: real measure)) (borel :: 'a::ordered_euclidean_space measure)" (is "p2e \ measurable ?P _") proof (safe intro!: borel_measurable_iff_halfspace_le[THEN iffD2]) - fix x i - let ?A = "{w \ space ?P. (p2e w :: 'a) $$ i \ x}" - assume "i < DIM('a)" - then have "?A = (\\<^isub>E j\{.. Basis" + then have "?A = (\\<^isub>E j\Basis. if i = j then {.. x} else UNIV)" using DIM_positive by (auto simp: space_PiM p2e_def PiE_def split: split_if_asm) then show "?A \ sets ?P" by auto qed lemma lborel_eq_lborel_space: - "(lborel :: 'a measure) = distr (\\<^isub>M i\{..\<^isub>M (i::'a::ordered_euclidean_space)\Basis. lborel) borel p2e" (is "?B = ?D") proof (rule lborel_eqI) show "sets ?D = sets borel" by simp - let ?P = "(\\<^isub>M i\{.. space ?P = (\\<^isub>E i\{.. space ?P = (\\<^isub>E i\Basis. {a \ i .. b \ i})" by (auto simp: eucl_le[where 'a='a] p2e_def space_PiM PiE_def Pi_iff) have "emeasure ?P (p2e -` {a..b} \ space ?P) = content {a..b}" proof cases assume "{a..b} \ {}" then have "a \ b" by (simp add: interval_ne_empty eucl_le[where 'a='a]) - then have "emeasure lborel {a..b} = (\xx\Basis. emeasure lborel {a \ x .. b \ x})" by (auto simp: content_closed_interval eucl_le[where 'a='a] intro!: setprod_ereal[symmetric]) also have "\ = emeasure ?P (p2e -` {a..b} \ space ?P)" @@ -986,13 +975,12 @@ lemma borel_fubini_positiv_integral: fixes f :: "'a::ordered_euclidean_space \ ereal" assumes f: "f \ borel_measurable borel" - shows "integral\<^isup>P lborel f = \\<^isup>+x. f (p2e x) \(\\<^isub>M i\{..P lborel f = \\<^isup>+x. f (p2e x) \(\\<^isub>M (i::'a)\Basis. lborel)" by (subst lborel_eq_lborel_space) (simp add: positive_integral_distr measurable_p2e f) lemma borel_fubini_integrable: fixes f :: "'a::ordered_euclidean_space \ real" - shows "integrable lborel f \ - integrable (\\<^isub>M i\{..x. f (p2e x))" + shows "integrable lborel f \ integrable (\\<^isub>M (i::'a)\Basis. lborel) (\x. f (p2e x))" (is "_ \ integrable ?B ?f") proof assume "integrable lborel f" @@ -1017,7 +1005,7 @@ lemma borel_fubini: fixes f :: "'a::ordered_euclidean_space \ real" assumes f: "f \ borel_measurable borel" - shows "integral\<^isup>L lborel f = \x. f (p2e x) \((\\<^isub>M i\{..L lborel f = \x. f (p2e x) \((\\<^isub>M (i::'a)\Basis. lborel))" using f by (simp add: borel_fubini_positiv_integral lebesgue_integral_def) lemma integrable_on_borel_integrable: @@ -1066,7 +1054,7 @@ using borel_measurable_continuous_on_open[of "{a <..< b }" f "\x. x" borel 0] by (auto intro!: measurable_If[where P="\x. x = a"] measurable_If[where P="\x. x = b"]) also have "?g = ?f" - using `a \ b` by (auto intro!: ext split: split_indicator) + using `a \ b` by (intro ext) (auto split: split_indicator) finally show "?f \ borel_measurable lborel" by simp qed diff -r 46be26e02456 -r 899c9c4e4a4c src/HOL/Probability/Sigma_Algebra.thy --- a/src/HOL/Probability/Sigma_Algebra.thy Fri Dec 14 14:46:01 2012 +0100 +++ b/src/HOL/Probability/Sigma_Algebra.thy Fri Dec 14 15:46:01 2012 +0100 @@ -322,6 +322,20 @@ ultimately show ?thesis by simp qed +lemma (in sigma_algebra) countable_UN': + fixes A :: "'i \ 'a set" + assumes X: "countable X" + assumes A: "A`X \ M" + shows "(\x\X. A x) \ M" +proof - + have "(\x\X. A x) = (\i\to_nat_on X ` X. A (from_nat_into X i))" + using X by auto + also have "\ \ M" + using A X + by (intro countable_UN) auto + finally show ?thesis . +qed + lemma (in sigma_algebra) countable_INT [intro]: fixes A :: "'i::countable \ 'a set" assumes A: "A`X \ M" "X \ {}" @@ -335,6 +349,20 @@ ultimately show ?thesis by metis qed +lemma (in sigma_algebra) countable_INT': + fixes A :: "'i \ 'a set" + assumes X: "countable X" "X \ {}" + assumes A: "A`X \ M" + shows "(\x\X. A x) \ M" +proof - + have "(\x\X. A x) = (\i\to_nat_on X ` X. A (from_nat_into X i))" + using X by auto + also have "\ \ M" + using A X + by (intro countable_INT) auto + finally show ?thesis . +qed + lemma ring_of_sets_Pow: "ring_of_sets sp (Pow sp)" by (auto simp: ring_of_sets_iff) @@ -365,6 +393,16 @@ with assms show ?thesis by auto qed +lemma (in sigma_algebra) sets_Collect_countable_Ex': + assumes "\i. {x\\. P i x} \ M" + assumes "countable I" + shows "{x\\. \i\I. P i x} \ M" +proof - + have "{x\\. \i\I. P i x} = (\i\I. {x\\. P i x})" by auto + with assms show ?thesis + by (auto intro!: countable_UN') +qed + lemmas (in sigma_algebra) sets_Collect = sets_Collect_imp sets_Collect_disj sets_Collect_conj sets_Collect_neg sets_Collect_const sets_Collect_countable_All sets_Collect_countable_Ex sets_Collect_countable_All