# HG changeset patch # User immler # Date 1452849271 -3600 # Node ID 9ca00b65d36c3b0e55ffd7308539d20016a38bce # Parent 4025b5ce19016b39cf0e3e38da5591706ffc344d continuity of parameterized integral; easier-to-apply formulation of rules diff -r 4025b5ce1901 -r 9ca00b65d36c src/HOL/Multivariate_Analysis/Integration.thy --- a/src/HOL/Multivariate_Analysis/Integration.thy Thu Jan 14 16:45:13 2016 +0100 +++ b/src/HOL/Multivariate_Analysis/Integration.thy Fri Jan 15 10:14:31 2016 +0100 @@ -11625,6 +11625,108 @@ by (force intro!: exI[where x="\C\CC. X0 (c C)"]) qed +lemma continuous_on_prod_compactE: + fixes fx::"'a::topological_space \ 'b::topological_space \ 'c::metric_space" + and e::real + assumes cont_fx: "continuous_on (U \ C) fx" + assumes "compact C" + assumes [intro]: "x0 \ U" + notes [continuous_intros] = continuous_on_compose2[OF cont_fx] + assumes "e > 0" + obtains X0 where "x0 \ X0" "open X0" + "\x\X0 \ U. \t \ C. dist (fx (x, t)) (fx (x0, t)) \ e" +proof - + def psi \ "\(x, t). dist (fx (x, t)) (fx (x0, t))" + def W0 \ "{(x, t) \ U \ C. psi (x, t) < e}" + have W0_eq: "W0 = psi -` {.. U \ C" + by (auto simp: vimage_def W0_def) + have "open {.. C) psi" + by (auto intro!: continuous_intros simp: psi_def split_beta') + from this[unfolded continuous_on_open_invariant, rule_format, OF \open {..] + obtain W where W: "open W" "W \ U \ C = W0 \ U \ C" + unfolding W0_eq by blast + have "{x0} \ C \ W \ U \ C" + unfolding W + by (auto simp: W0_def psi_def \0 < e\) + then have "{x0} \ C \ W" by blast + from tube_lemma[OF \compact C\ \open W\ this] + obtain X0 where X0: "x0 \ X0" "open X0" "X0 \ C \ W" + by blast + + have "\x\X0 \ U. \t \ C. dist (fx (x, t)) (fx (x0, t)) \ e" + proof safe + fix x assume x: "x \ X0" "x \ U" + fix t assume t: "t \ C" + have "dist (fx (x, t)) (fx (x0, t)) = psi (x, t)" + by (auto simp: psi_def) + also + { + have "(x, t) \ X0 \ C" + using t x + by auto + also note \\ \ W\ + finally have "(x, t) \ W" . + with t x have "(x, t) \ W \ U \ C" + by blast + also note \W \ U \ C = W0 \ U \ C\ + finally have "psi (x, t) < e" + by (auto simp: W0_def) + } + finally show "dist (fx (x, t)) (fx (x0, t)) \ e" by simp + qed + from X0(1,2) this show ?thesis .. +qed + +lemma integral_continuous_on_param: + fixes f::"'a::topological_space \ 'b::euclidean_space \ 'c::banach" + assumes cont_fx: "continuous_on (U \ cbox a b) (\(x, t). f x t)" + shows "continuous_on U (\x. integral (cbox a b) (f x))" +proof cases + assume "content (cbox a b) \ 0" + then have ne: "cbox a b \ {}" by auto + + note [continuous_intros] = + continuous_on_compose2[OF cont_fx, where f="\y. Pair x y" for x, + unfolded split_beta fst_conv snd_conv] + + show ?thesis + unfolding continuous_on_def + proof (safe intro!: tendstoI) + fix e'::real and x + assume "e' > 0" + def e \ "e' / (content (cbox a b) + 1)" + have "e > 0" using \e' > 0\ by (auto simp: e_def intro!: divide_pos_pos add_nonneg_pos) + assume "x \ U" + from continuous_on_prod_compactE[OF cont_fx compact_cbox \x \ U\ \0 < e\] + obtain X0 where X0: "x \ X0" "open X0" + and fx_bound: "\y t. y \ X0 \ U \ t \ cbox a b \ norm (f y t - f x t) \ e" + unfolding split_beta fst_conv snd_conv dist_norm + by metis + have "\\<^sub>F y in at x within U. y \ X0 \ U" + using X0(1) X0(2) eventually_at_topological by auto + then show "\\<^sub>F y in at x within U. dist (integral (cbox a b) (f y)) (integral (cbox a b) (f x)) < e'" + proof eventually_elim + case (elim y) + have "dist (integral (cbox a b) (f y)) (integral (cbox a b) (f x)) = + norm (integral (cbox a b) (\t. f y t - f x t))" + using elim \x \ U\ + unfolding dist_norm + by (subst integral_diff) + (auto intro!: integrable_continuous continuous_intros) + also have "\ \ e * content (cbox a b)" + using elim \x \ U\ + by (intro integrable_bound) + (auto intro!: fx_bound \x \ U \ less_imp_le[OF \0 < e\] + integrable_continuous continuous_intros) + also have "\ < e'" + using \0 < e'\ \e > 0\ + by (auto simp: e_def divide_simps) + finally show "dist (integral (cbox a b) (f y)) (integral (cbox a b) (f x)) < e'" . + qed + qed +qed (auto intro!: continuous_on_const) + lemma eventually_closed_segment: fixes x0::"'a::real_normed_vector" assumes "open X0" "x0 \ X0" @@ -11646,47 +11748,36 @@ qed lemma leibniz_rule: - fixes f::"'a::banach \ 'b::euclidean_space \ 'c::banach" - assumes fx: "\x t. x \ U \ t \ (cbox a b) \ - ((\x. f (x, t)) has_derivative blinfun_apply (fx (x, t))) (at x within U)" - assumes integrable_f2: "\x. x \ U \ (\t. f (x, t)) integrable_on cbox a b" - assumes cont_fx: "continuous_on (U \ (cbox a b)) fx" + fixes f::"'a::banach \ 'b::euclidean_space \ 'c::banach" + assumes fx: "\x t. x \ U \ t \ cbox a b \ + ((\x. f x t) has_derivative blinfun_apply (fx x t)) (at x within U)" + assumes integrable_f2: "\x. x \ U \ f x integrable_on cbox a b" + assumes cont_fx: "continuous_on (U \ (cbox a b)) (\(x, t). fx x t)" assumes [intro]: "x0 \ U" assumes "convex U" - notes [continuous_intros] = continuous_on_compose2[OF cont_fx] shows - "((\x. integral (cbox a b) (\t. f (x, t))) has_derivative integral (cbox a b) (\t. fx (x0, t))) - (at x0 within U)" + "((\x. integral (cbox a b) (f x)) has_derivative integral (cbox a b) (fx x0)) (at x0 within U)" (is "(?F has_derivative ?dF) _") proof cases assume "content (cbox a b) \ 0" then have ne: "cbox a b \ {}" by auto + note [continuous_intros] = + continuous_on_compose2[OF cont_fx, where f="\y. Pair x y" for x, + unfolded split_beta fst_conv snd_conv] show ?thesis proof (intro has_derivativeI bounded_linear_scaleR_left tendstoI, fold norm_conv_dist) - have cont_f1: "\t. t \ cbox a b \ continuous_on U (\x. f (x, t))" + have cont_f1: "\t. t \ cbox a b \ continuous_on U (\x. f x t)" by (auto simp: continuous_on_eq_continuous_within intro!: has_derivative_continuous fx) note [continuous_intros] = continuous_on_compose2[OF cont_f1] fix e'::real assume "e' > 0" def e \ "e' / (content (cbox a b) + 1)" have "e > 0" using \e' > 0\ by (auto simp: e_def intro!: divide_pos_pos add_nonneg_pos) - def psi \ "\(x, t). fx (x, t) - fx (x0, t)" - def W0 \ "{(x, t) \ U \ (cbox a b). norm (psi (x, t)) < e}" - have W0_eq: "W0 = (\(x, t). norm (psi (x, t))) -` {.. U \ (cbox a b)" - by (auto simp: vimage_def W0_def) - have "open {.. (cbox a b)) (\(x, t). norm (psi (x, t)))" - by (auto intro!: continuous_intros simp: psi_def split_beta') - from this[unfolded continuous_on_open_invariant, rule_format, OF \open {..] - obtain W where W: "open W" "W \ U \ (cbox a b) = W0 \ U \ (cbox a b)" - unfolding W0_eq by blast - have "{x0} \ (cbox a b) \ W \ U \ (cbox a b)" - unfolding W - by (auto simp: W0_def psi_def \0 < e\) - then have "{x0} \ cbox a b \ W" by blast - from tube_lemma[OF compact_cbox[of a b] \open W\ this] - obtain X0 where X0: "x0 \ X0" "open X0" "X0 \ cbox a b \ W" - by blast + from continuous_on_prod_compactE[OF cont_fx compact_cbox \x0 \ U\ \e > 0\] + obtain X0 where X0: "x0 \ X0" "open X0" + and fx_bound: "\x t. x \ X0 \ U \ t \ cbox a b \ norm (fx x t - fx x0 t) \ e" + unfolding split_beta fst_conv snd_conv + by (metis dist_norm) note eventually_closed_segment[OF \open X0\ \x0 \ X0\, of U] moreover @@ -11706,7 +11797,7 @@ from elim have [intro]: "x \ U" by auto have "?F x - ?F x0 - ?dF (x - x0) = - integral (cbox a b) (\xa. f (x, xa) - f (x0, xa) - fx (x0, xa) (x - x0))" + integral (cbox a b) (\y. f x y - f x0 y - fx x0 y (x - x0))" (is "_ = ?id") using \x \ x0\ by (subst blinfun_apply_integral integral_diff, @@ -11715,42 +11806,21 @@ also { fix t assume t: "t \ (cbox a b)" - have seg: "\t. t \ {0..1} \ x0 + t *\<^sub>R (x - x0) \ closed_segment x0 x \ U" + have seg: "\t. t \ {0..1} \ x0 + t *\<^sub>R (x - x0) \ X0 \ U" using \closed_segment x0 x \ U\ + \closed_segment x0 x \ X0\ by (force simp: closed_segment_def algebra_simps) from t have deriv: - "((\x. f (x, t)) has_derivative (fx (y, t))) (at y within closed_segment x0 x \ U)" - if "y \ closed_segment x0 x \ U" for y + "((\x. f x t) has_derivative (fx y t)) (at y within X0 \ U)" + if "y \ X0 \ U" for y unfolding has_vector_derivative_def[symmetric] using that \x \ X0\ by (intro has_derivative_within_subset[OF fx]) auto - have "\x\closed_segment x0 x \ U. norm (fx (x, t) - fx (x0, t)) \ e" - proof - fix y assume y: "y \ closed_segment x0 x \ U" - have "norm (fx (y, t) - fx (x0, t)) = norm (psi (y, t))" - by (auto simp: psi_def) - also - { - have "(y, t) \ X0 \ (cbox a b)" - using t \closed_segment x0 x \ X0\ y - by auto - also note \\ \ W\ - finally have "(y, t) \ W" . - with t y have "(y, t) \ W \ U \ (cbox a b)" - by blast - also note \W \ U \ (cbox a b) = W0 \ U \ (cbox a b)\ - finally have "norm (psi (y, t)) < e" - by (auto simp: W0_def) - } - finally show "norm (fx (y, t) - fx (x0, t)) \ e" by simp - qed - then have onorm: - "\x\closed_segment x0 x \ U. - onorm (blinfun_apply (fx (x, t)) - blinfun_apply (fx (x0, t))) \ e" - by (auto simp: norm_blinfun.rep_eq blinfun.diff_left[abs_def] fun_diff_def) - - from differentiable_bound_linearization[OF seg deriv onorm] - have "norm (f (x, t) - f (x0, t) - fx (x0, t) (x - x0)) \ e * norm (x - x0)" + have "\x \ X0 \ U. onorm (blinfun_apply (fx x t) - (fx x0 t)) \ e" + using fx_bound t + by (auto simp add: norm_blinfun_def fun_diff_def blinfun.bilinear_simps[symmetric]) + from differentiable_bound_linearization[OF seg deriv this] X0 + have "norm (f x t - f x0 t - fx x0 t (x - x0)) \ e * norm (x - x0)" by (auto simp add: ac_simps) } then have "norm ?id \ integral (cbox a b) (\_. e * norm (x - x0))" @@ -11777,28 +11847,29 @@ by (simp add: has_vector_derivative_def) lemma leibniz_rule_vector_derivative: - fixes f::"real \ 'b::euclidean_space \ 'c::banach" - assumes fx: "\x t. x \ U \ t \ (cbox a b) \ - ((\x. f (x, t)) has_vector_derivative (fx (x, t))) (at x within U)" - assumes integrable_f2: "\x. x \ U \ (\t. f (x, t)) integrable_on cbox a b" - assumes cont_fx: "continuous_on (U \ (cbox a b)) fx" + fixes f::"real \ 'b::euclidean_space \ 'c::banach" + assumes fx: "\x t. x \ U \ t \ cbox a b \ + ((\x. f x t) has_vector_derivative (fx x t)) (at x within U)" + assumes integrable_f2: "\x. x \ U \ (f x) integrable_on cbox a b" + assumes cont_fx: "continuous_on (U \ cbox a b) (\(x, t). fx x t)" assumes U: "x0 \ U" "convex U" - notes [continuous_intros] = continuous_on_compose2[OF cont_fx] - shows - "((\x. integral (cbox a b) (\t. f (x, t))) has_vector_derivative - integral (cbox a b) (\t. fx (x0, t))) (at x0 within U)" -proof - - have *: "blinfun_scaleR_left (integral (cbox a b) (\t. fx (x0, t))) = - integral (cbox a b) (\t. blinfun_scaleR_left (fx (x0, t)))" + shows "((\x. integral (cbox a b) (f x)) has_vector_derivative integral (cbox a b) (fx x0)) + (at x0 within U)" +proof - + note [continuous_intros] = + continuous_on_compose2[OF cont_fx, where f="\y. Pair x y" for x, + unfolded split_beta fst_conv snd_conv] + have *: "blinfun_scaleR_left (integral (cbox a b) (fx x0)) = + integral (cbox a b) (\t. blinfun_scaleR_left (fx x0 t))" by (subst integral_linear[symmetric]) (auto simp: has_vector_derivative_def o_def intro!: integrable_continuous U continuous_intros bounded_linear_intros) show ?thesis unfolding has_vector_derivative_eq_has_derivative_blinfun apply (rule has_derivative_eq_rhs) - apply (rule leibniz_rule[OF _ integrable_f2 _ U, where fx="\x. blinfun_scaleR_left (fx x)"]) - using assms(1) - apply (auto simp: has_vector_derivative_def * intro!: continuous_intros) + apply (rule leibniz_rule[OF _ integrable_f2 _ U, where fx="\x t. blinfun_scaleR_left (fx x t)"]) + using fx cont_fx + apply (auto simp: has_vector_derivative_def * split_beta intro!: continuous_intros) done qed @@ -11808,25 +11879,27 @@ by (simp add: has_field_derivative_def) lemma leibniz_rule_field_derivative: - fixes f::"'a::{real_normed_field, banach} \ 'b::euclidean_space \ 'a" - assumes fx: "\x t. x \ U \ t \ (cbox a b) \ ((\x. f (x, t)) has_field_derivative (fx (x, t))) (at x within U)" - assumes integrable_f2: "\x. x \ U \ (\t. f (x, t)) integrable_on cbox a b" - assumes cont_fx: "continuous_on (U \ (cbox a b)) fx" + fixes f::"'a::{real_normed_field, banach} \ 'b::euclidean_space \ 'a" + assumes fx: "\x t. x \ U \ t \ cbox a b \ ((\x. f x t) has_field_derivative fx x t) (at x within U)" + assumes integrable_f2: "\x. x \ U \ (f x) integrable_on cbox a b" + assumes cont_fx: "continuous_on (U \ (cbox a b)) (\(x, t). fx x t)" assumes U: "x0 \ U" "convex U" - notes [continuous_intros] = continuous_on_compose2[OF cont_fx] - shows "((\x. integral (cbox a b) (\t. f (x, t))) has_field_derivative integral (cbox a b) (\t. fx (x0, t))) (at x0 within U)" -proof - - have *: "blinfun_mult_right (integral (cbox a b) (\t. fx (x0, t))) = - integral (cbox a b) (\t. blinfun_mult_right (fx (x0, t)))" + shows "((\x. integral (cbox a b) (f x)) has_field_derivative integral (cbox a b) (fx x0)) (at x0 within U)" +proof - + note [continuous_intros] = + continuous_on_compose2[OF cont_fx, where f="\y. Pair x y" for x, + unfolded split_beta fst_conv snd_conv] + have *: "blinfun_mult_right (integral (cbox a b) (fx x0)) = + integral (cbox a b) (\t. blinfun_mult_right (fx x0 t))" by (subst integral_linear[symmetric]) (auto simp: has_vector_derivative_def o_def intro!: integrable_continuous U continuous_intros bounded_linear_intros) show ?thesis unfolding has_field_derivative_eq_has_derivative_blinfun apply (rule has_derivative_eq_rhs) - apply (rule leibniz_rule[OF _ integrable_f2 _ U, where fx="\x. blinfun_mult_right (fx x)"]) - using assms(1) - apply (auto simp: has_field_derivative_def * intro!: continuous_intros) + apply (rule leibniz_rule[OF _ integrable_f2 _ U, where fx="\x t. blinfun_mult_right (fx x t)"]) + using fx cont_fx + apply (auto simp: has_field_derivative_def * split_beta intro!: continuous_intros) done qed