diff -r 46be26e02456 -r 899c9c4e4a4c src/HOL/Multivariate_Analysis/Integration.thy --- a/src/HOL/Multivariate_Analysis/Integration.thy Fri Dec 14 14:46:01 2012 +0100 +++ b/src/HOL/Multivariate_Analysis/Integration.thy Fri Dec 14 15:46:01 2012 +0100 @@ -71,12 +71,8 @@ apply (auto simp: isUb_def setle_def) done -lemma bounded_linear_component [intro]: "bounded_linear (\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