# HG changeset patch # User paulson # Date 1506617935 -3600 # Node ID 41bf4d324ac416542574adc4937f8447a943dc75 # Parent 0b9e6ce3b843372e636b25752d0c240164d66215# Parent bd7901f702c9a7adac6aa3442c72fb81a64404de merged diff -r 0b9e6ce3b843 -r 41bf4d324ac4 src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy --- a/src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy Thu Sep 28 09:42:28 2017 +0200 +++ b/src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy Thu Sep 28 17:58:55 2017 +0100 @@ -2253,9 +2253,7 @@ have "(\x. \f x \ i\) integrable_on S" using assms integrable_component [OF fcomp, where y=i] that by simp then have "(\x. f x \ i) absolutely_integrable_on S" - apply - - apply (rule abs_absolutely_integrableI_1, auto) - by (simp add: f integrable_component) + using abs_absolutely_integrableI_1 f integrable_component by blast then show ?thesis by (rule absolutely_integrable_scaleR_right) qed diff -r 0b9e6ce3b843 -r 41bf4d324ac4 src/HOL/Analysis/Henstock_Kurzweil_Integration.thy --- a/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy Thu Sep 28 09:42:28 2017 +0200 +++ b/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy Thu Sep 28 17:58:55 2017 +0100 @@ -6049,8 +6049,7 @@ have "norm (integral S f) \ integral S ((\x. x \ k) \ g)" apply (rule integral_norm_bound_integral[OF f integrable_linear[OF g]]) apply (simp add: bounded_linear_inner_left) - unfolding o_def - apply (metis fg) + apply (metis fg o_def) done then show ?thesis unfolding o_def integral_component_eq[OF g] . @@ -6167,7 +6166,6 @@ have "closed_segment x0 x \ U" by (rule \convex U\[unfolded convex_contains_segment, rule_format, OF \x0 \ U\ \x \ U\]) from elim have [intro]: "x \ U" by auto - have "?F x - ?F x0 - ?dF (x - x0) = integral (cbox a b) (\y. f x y - f x0 y - fx x0 y (x - x0))" (is "_ = ?id") @@ -6204,7 +6202,7 @@ also have "\ < e' * norm (x - x0)" using \e' > 0\ apply (intro mult_strict_right_mono[OF _ \0 < norm (x - x0)\]) - apply (auto simp: divide_simps e_def) + apply (auto simp: divide_simps e_def) by (metis \0 < e\ e_def order.asym zero_less_divide_iff) finally have "norm (?F x - ?F x0 - ?dF (x - x0)) < e' * norm (x - x0)" . then show ?case @@ -6293,14 +6291,12 @@ by atomize_elim (auto simp: integrable_on_def intro!: choice) moreover - have gi[simp]: "g integrable_on (cbox a b)" by (auto intro!: integrable_continuous uniform_limit_theorem[OF _ u] eventuallyI c) then obtain J where J: "(g has_integral J) (cbox a b)" by blast moreover - have "(I \ J) F" proof cases assume "content (cbox a b) = 0" @@ -6449,20 +6445,17 @@ f integrable_continuous_real)+ have deriv: "(((\x. integral {c..x} f) \ g) has_vector_derivative g' x *\<^sub>R f (g x)) (at x within {a..b})" if "x \ {a..b} - s" for x - apply (rule has_vector_derivative_eq_rhs) - apply (rule vector_diff_chain_within) - apply (subst has_field_derivative_iff_has_vector_derivative [symmetric]) - apply (rule deriv that)+ - apply (rule has_vector_derivative_within_subset) - apply (rule integral_has_vector_derivative f)+ - using that le subset - apply blast+ - done + proof (rule has_vector_derivative_eq_rhs [OF vector_diff_chain_within refl]) + show "(g has_vector_derivative g' x) (at x within {a..b})" + using deriv has_field_derivative_iff_has_vector_derivative that by blast + show "((\x. integral {c..x} f) has_vector_derivative f (g x)) + (at (g x) within g ` {a..b})" + using that le subset + by (blast intro: has_vector_derivative_within_subset integral_has_vector_derivative f) + qed have deriv: "(?F has_vector_derivative g' x *\<^sub>R f (g x)) (at x)" if "x \ {a..b} - (s \ {a,b})" for x using deriv[of x] that by (simp add: at_within_closed_interval o_def) - - have "((\x. g' x *\<^sub>R f (g x)) has_integral (?F b - ?F a)) {a..b}" using le cont_int s deriv cont_int by (intro fundamental_theorem_of_calculus_interior_strong[of "s \ {a,b}"]) simp_all @@ -6794,20 +6787,21 @@ \ e * content (cbox (u,w) (v,z)) / content ?CBOX" by (rule norm_xx [OF integral_Pair_const 1 2]) } note * = this - show "norm (integral ?CBOX f - integral (cbox a b) (\x. integral (cbox c d) (\y. f (x,y)))) \ e" + have "norm (integral ?CBOX f - integral (cbox a b) (\x. integral (cbox c d) (\y. f (x,y)))) \ e" + if "\x\?CBOX. \x'\?CBOX. norm (x' - x) < k \ norm (f x' - f x) < e /(2 * content (?CBOX))" "0 < k" for k + proof - + obtain p where ptag: "p tagged_division_of cbox (a, c) (b, d)" + and fine: "(\x. ball x k) fine p" + using fine_division_exists \0 < k\ by blast + show ?thesis + apply (rule op_acbd [OF division_of_tagged_division [OF ptag]]) + using that fine ptag \0 < k\ by (auto simp: *) + qed + then show "norm (integral ?CBOX f - integral (cbox a b) (\x. integral (cbox c d) (\y. f (x,y)))) \ e" using compact_uniformly_continuous [OF assms compact_cbox] apply (simp add: uniformly_continuous_on_def dist_norm) apply (drule_tac x="e/2 / content?CBOX" in spec) - using cbp \0 < e\ - apply (auto simp: zero_less_mult_iff) - apply (rename_tac k) - apply (rule_tac e1=k in fine_division_exists [OF gauge_ball, where a = "(a,c)" and b = "(b,d)"]) - apply assumption - apply (rule op_acbd) - apply (erule division_of_tagged_division) - using * - apply auto - done + using cbp \0 < e\ by (auto simp: zero_less_mult_iff) qed then show ?thesis by simp @@ -6850,7 +6844,6 @@ shows "((\x::real. exp (-a*x)) has_integral exp (-a*c)/a) {c..}" proof - define f where "f = (\k x. if x \ {c..real k} then exp (-a*x) else 0)" - { fix k :: nat assume k: "of_nat k \ c" from k a diff -r 0b9e6ce3b843 -r 41bf4d324ac4 src/HOL/Analysis/Tagged_Division.thy --- a/src/HOL/Analysis/Tagged_Division.thy Thu Sep 28 09:42:28 2017 +0200 +++ b/src/HOL/Analysis/Tagged_Division.thy Thu Sep 28 17:58:55 2017 +0100 @@ -1919,89 +1919,71 @@ lemma interval_bisection_step: fixes type :: "'a::euclidean_space" - assumes "P {}" - and "\s t. P s \ P t \ interior(s) \ interior(t) = {} \ P (s \ t)" - and "\ P (cbox a (b::'a))" + assumes emp: "P {}" + and Un: "\S T. \P S; P T; interior(S) \ interior(T) = {}\ \ P (S \ T)" + and non: "\ P (cbox a (b::'a))" obtains c d where "\ P (cbox c d)" - and "\i\Basis. a\i \ c\i \ c\i \ d\i \ d\i \ b\i \ 2 * (d\i - c\i) \ b\i - a\i" + and "\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 "cbox a b \ {}" - using assms(1,3) by metis + using emp non by metis then have ab: "\i. i\Basis \ a \ i \ b \ i" by (force simp: mem_box) - have UN_cases: "\finite f; - \s. s\f \ P s; - \s. s\f \ \a b. s = cbox a b; - \s t. s\f \ t\f \ s \ t \ interior s \ interior t = {}\ \ P (\f)" for f - proof (induct f rule: finite_induct) - case empty - show ?case - using assms(1) by auto + have UN_cases: "\finite \; + \S. S\\ \ P S; + \S. S\\ \ \a b. S = cbox a b; + \S T. S\\ \ T\\ \ S \ T \ interior S \ interior T = {}\ \ P (\\)" for \ + proof (induct \ rule: finite_induct) + case empty show ?case + using emp by auto next case (insert x f) - show ?case - unfolding Union_insert - apply (rule assms(2)[rule_format]) - using Int_interior_Union_intervals [of f "interior x"] - by (metis (no_types, lifting) insert insert_iff open_interior) + then show ?case + unfolding Union_insert by (metis Int_interior_Union_intervals Un insert_iff open_interior) qed - let ?A = "{cbox 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 (cbox c d) \ False" - then show thesis - unfolding atomize_not not_all - by (blast intro: that) - } - assume as: "\c d. ?PP c d \ P (cbox c d)" - have "P (\?A)" + let ?ab = "\i. (a\i + b\i) / 2" + let ?A = "{cbox c d | c d::'a. \i\Basis. (c\i = a\i) \ (d\i = ?ab i) \ + (c\i = ?ab i) \ (d\i = b\i)}" + have "P (\?A)" + if "\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 \ P (cbox c d)" proof (rule UN_cases) - let ?B = "(\s. cbox (\i\Basis. (if i \ s then a\i else (a\i + b\i) / 2) *\<^sub>R i::'a) - (\i\Basis. (if i \ s then (a\i + b\i) / 2 else b\i) *\<^sub>R i)) ` {s. s \ Basis}" + let ?B = "(\S. cbox (\i\Basis. (if i \ S then a\i else ?ab i) *\<^sub>R i::'a) + (\i\Basis. (if i \ S then ?ab i else b\i) *\<^sub>R i)) ` {s. s \ Basis}" have "?A \ ?B" proof fix x assume "x \ ?A" then obtain c d where x: "x = cbox c d" - "\i. i \ Basis \ - c \ i = a \ i \ d \ i = (a \ i + b \ i) / 2 \ - c \ i = (a \ i + b \ i) / 2 \ d \ i = b \ i" by blast - show "x \ ?B" - unfolding image_iff x - apply (rule_tac x="{i. i\Basis \ c\i = a\i}" in bexI) - apply (rule arg_cong2 [where f = cbox]) - using x(2) ab - apply (auto simp add: euclidean_eq_iff[where 'a='a]) - by fastforce + "\i. i \ Basis \ + c \ i = a \ i \ d \ i = ?ab i \ c \ i = ?ab i \ d \ i = b \ i" + by blast + have "c = (\i\Basis. (if c \ i = a \ i then a \ i else ?ab i) *\<^sub>R i)" + "d = (\i\Basis. (if c \ i = a \ i then ?ab i else b \ i) *\<^sub>R i)" + using x(2) ab by (fastforce simp add: euclidean_eq_iff[where 'a='a])+ + then show "x \ ?B" + unfolding x by (rule_tac x="{i. i\Basis \ c\i = a\i}" in image_eqI) auto qed then show "finite ?A" by (rule finite_subset) auto next - fix s - assume "s \ ?A" + fix S + assume "S \ ?A" then obtain c d - where s: "s = cbox c d" - "\i. i \ Basis \ - c \ i = a \ i \ d \ i = (a \ i + b \ i) / 2 \ - c \ i = (a \ i + b \ i) / 2 \ d \ i = b \ i" + where s: "S = cbox c d" + "\i. i \ Basis \ c \ i = a \ i \ d \ i = ?ab i \ c \ i = ?ab i \ d \ i = b \ i" by blast - show "P s" - unfolding s - apply (rule as[rule_format]) - using ab s(2) by force - show "\a b. s = cbox a b" + show "P S" + unfolding s using ab s(2) by (fastforce intro!: that) + show "\a b. S = cbox a b" unfolding s by auto - fix t - assume "t \ ?A" + fix T + assume "T \ ?A" then obtain e f where t: - "t = cbox e f" - "\i. i \ Basis \ - e \ i = a \ i \ f \ i = (a \ i + b \ i) / 2 \ - e \ i = (a \ i + b \ i) / 2 \ f \ i = b \ i" + "T = cbox e f" + "\i. i \ Basis \ e \ i = a \ i \ f \ i = ?ab i \ e \ i = ?ab i \ f \ i = b \ i" by blast - assume "s \ t" + assume "S \ T" then have "\ (c = e \ d = f)" unfolding s t by auto then obtain i where "c\i \ e\i \ d\i \ f\i" and i': "i \ Basis" @@ -2011,24 +1993,15 @@ using t(2)[OF i'] \c \ i \ e \ i \ d \ i \ f \ i\ i' s(2) t(2) by fastforce have *: "\s t. (\a. a \ s \ a \ t \ False) \ s \ t = {}" by auto - show "interior s \ interior t = {}" + show "interior S \ interior T = {}" unfolding s t interior_cbox proof (rule *) fix x assume "x \ box c d" "x \ box e f" then have x: "c\i < d\i" "e\i < f\i" "c\i < f\i" "e\i < d\i" - unfolding mem_box using i' - by force+ - show False using s(2)[OF i'] - proof safe - assume as: "c \ i = a \ i" "d \ i = (a \ i + b \ i) / 2" - show False - using t(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 t(2)[OF i'] and i x unfolding as by(fastforce simp add:field_simps) - qed + unfolding mem_box using i' by force+ + show False using s(2)[OF i'] t(2)[OF i'] and i x + by auto qed qed also have "\?A = cbox a b" @@ -2037,48 +2010,30 @@ assume "x \ \?A" then obtain c d where x: "x \ cbox c d" - "\i. i \ Basis \ - c \ i = a \ i \ d \ i = (a \ i + b \ i) / 2 \ - c \ i = (a \ i + b \ i) / 2 \ d \ i = b \ i" + "\i. i \ Basis \ c \ i = a \ i \ d \ i = ?ab i \ c \ i = ?ab i \ d \ i = b \ i" by blast - show "x\cbox a b" - unfolding mem_box - proof safe - fix i :: 'a - assume i: "i \ Basis" - then show "a \ i \ x \ i" "x \ i \ b \ i" - using x(2)[OF i] x(1)[unfolded mem_box,THEN bspec, OF i] by auto - qed + then show "x\cbox a b" + unfolding mem_box by force next fix x assume x: "x \ cbox a b" - 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" + then have "\i\Basis. \c d. (c = a\i \ d = ?ab i \ c = ?ab i \ d = b\i) \ c\x\i \ x\i \ d" (is "\i\Basis. \c d. ?P i c d") - unfolding mem_box - 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_box,THEN bspec, OF i] by auto - then show "\c d. ?P i c d" - by blast - qed - then obtain \ \ where - "\i\Basis. (\ \ i = a \ i \ \ \ i = (a \ i + b \ i) / 2 \ - \ \ i = (a \ i + b \ i) / 2 \ \ \ i = b \ i) \ \ \ i \ x \ i \ x \ i \ \ \ i" + unfolding mem_box by (metis linear) + then obtain \ \ where "\i\Basis. (\ \ i = a \ i \ \ \ i = ?ab i \ + \ \ i = ?ab i \ \ \ i = b \ i) \ \ \ i \ x \ i \ x \ i \ \ \ i" by (auto simp: choice_Basis_iff) then show "x\\?A" by (force simp add: mem_box) qed - finally show False - using assms by auto + finally show thesis + by (metis (no_types, lifting) assms(3) that) qed lemma interval_bisection: fixes type :: "'a::euclidean_space" assumes "P {}" - and "(\s t. P s \ P t \ interior(s) \ interior(t) = {} \ P(s \ t))" + and Un: "\S T. \P S; P T; interior(S) \ interior(T) = {}\ \ P (S \ T)" and "\ P (cbox a (b::'a))" obtains x where "x \ cbox a b" and "\e>0. \c d. x \ cbox c d \ cbox c d \ ball x e \ cbox c d \ cbox a b \ \ P (cbox c d)" @@ -2092,14 +2047,14 @@ case True then show ?thesis by auto next - case as: False + case False obtain c d where "\ P (cbox c d)" - "\i\Basis. + "\i. i \ Basis \ fst x \ i \ c \ i \ c \ i \ d \ i \ d \ i \ snd x \ i \ 2 * (d \ i - c \ i) \ snd x \ i - fst x \ i" - by (rule interval_bisection_step[of P, OF assms(1-2) as]) + by (blast intro: interval_bisection_step[of P, OF assms(1-2) False]) then show ?thesis by (rule_tac x="(c,d)" in exI) auto qed @@ -2281,33 +2236,17 @@ lemma tagged_division_finer: fixes p :: "('a::euclidean_space \ ('a::euclidean_space set)) set" - assumes "p tagged_division_of (cbox a b)" + assumes ptag: "p tagged_division_of (cbox a b)" and "gauge d" obtains q where "q tagged_division_of (cbox a b)" and "d fine q" and "\(x,k) \ p. k \ d(x) \ (x,k) \ q" proof - - let ?P = "\p. p tagged_partial_division_of (cbox a b) \ gauge d \ - (\q. q tagged_division_of (\{k. \x. (x,k) \ p}) \ d fine q \ - (\(x,k) \ p. k \ d(x) \ (x,k) \ q))" - { - have *: "finite p" "p tagged_partial_division_of (cbox a b)" - using assms(1) - unfolding tagged_division_of_def - by auto - presume "\p. finite p \ ?P p" - from this[rule_format,OF * assms(2)] - obtain q where q: "q tagged_division_of \{k. \x. (x, k) \ p}" "d fine q" "(\(x, k)\p. k \ d x \ (x, k) \ q)" - by auto - with that[of q] show ?thesis - using assms(1) by auto - } - fix p :: "('a::euclidean_space \ ('a::euclidean_space set)) set" - assume as: "finite p" - show "?P p" - apply rule - apply rule - using as + have p: "finite p" "p tagged_partial_division_of (cbox a b)" + using ptag unfolding tagged_division_of_def by auto + have "(\q. q tagged_division_of (\{k. \x. (x,k) \ p}) \ d fine q \ (\(x,k) \ p. k \ d(x) \ (x,k) \ q))" + if "finite p" "p tagged_partial_division_of (cbox a b)" "gauge d" for p + using that proof (induct p) case empty show ?case @@ -2325,7 +2264,7 @@ unfolding xk by auto note p = tagged_partial_division_ofD[OF insert(4)] obtain u v where uv: "k = cbox u v" - using p(4)[unfolded xk, OF insertI1] by blast + using p(4) xk by blast have "finite {k. \x. (x, k) \ p}" apply (rule finite_subset[of _ "snd ` p"]) using image_iff apply fastforce @@ -2363,6 +2302,9 @@ done qed qed + with p obtain q where q: "q tagged_division_of \{k. \x. (x, k) \ p}" "d fine q" "\(x, k)\p. k \ d x \ (x, k) \ q" + by (meson \gauge d\) + with ptag that show ?thesis by auto qed subsubsection \Covering lemma\