# HG changeset patch # User paulson # Date 1442861533 -3600 # Node ID 3e491e34a62e23edc0a865197d5ba751d33d99a4 # Parent a8a8eca8580177e31fb221d8b2c9e7cf15a88463 new lemmas and movement of lemmas into place diff -r a8a8eca85801 -r 3e491e34a62e src/HOL/Deriv.thy --- a/src/HOL/Deriv.thy Mon Sep 21 14:44:32 2015 +0200 +++ b/src/HOL/Deriv.thy Mon Sep 21 19:52:13 2015 +0100 @@ -644,6 +644,18 @@ ((\x. f x - g x) has_vector_derivative (f' - g')) net" by (auto simp: has_vector_derivative_def scaleR_diff_right) +lemma has_vector_derivative_add_const: + "((\t. g t + z) has_vector_derivative f') (at x) = ((\t. g t) has_vector_derivative f') (at x)" +apply (intro iffI) +apply (drule has_vector_derivative_diff [where g = "\t. z", OF _ has_vector_derivative_const], simp) +apply (drule has_vector_derivative_add [OF _ has_vector_derivative_const], simp) +done + +lemma has_vector_derivative_diff_const: + "((\t. g t - z) has_vector_derivative f') (at x) = ((\t. g t) has_vector_derivative f') (at x)" +using has_vector_derivative_add_const [where z = "-z"] +by simp + lemma (in bounded_linear) has_vector_derivative: assumes "(g has_vector_derivative g') F" shows "((\x. f (g x)) has_vector_derivative f g') F" diff -r a8a8eca85801 -r 3e491e34a62e src/HOL/Fun.thy --- a/src/HOL/Fun.thy Mon Sep 21 14:44:32 2015 +0200 +++ b/src/HOL/Fun.thy Mon Sep 21 19:52:13 2015 +0100 @@ -69,7 +69,7 @@ lemma comp_eq_elim: "a o b = c o d \ ((\v. a (b v) = c (d v)) \ R) \ R" - by (simp add: fun_eq_iff) + by (simp add: fun_eq_iff) lemma comp_eq_dest_lhs: "a o b = c \ a (b v) = c v" by clarsimp @@ -833,7 +833,7 @@ thus False by best qed -subsection \Setup\ +subsection \Setup\ subsubsection \Proof tools\ diff -r a8a8eca85801 -r 3e491e34a62e src/HOL/Int.thy --- a/src/HOL/Int.thy Mon Sep 21 14:44:32 2015 +0200 +++ b/src/HOL/Int.thy Mon Sep 21 19:52:13 2015 +0100 @@ -518,7 +518,7 @@ lemma int_cases3 [case_names zero pos neg]: fixes k :: int assumes "k = 0 \ P" and "\n. k = int n \ n > 0 \ P" - and "\n. k = - int n \ n > 0 \ P" + and "\n. k = - int n \ n > 0 \ P" shows "P" proof (cases k "0::int" rule: linorder_cases) case equal with assms(1) show P by simp @@ -890,7 +890,7 @@ moreover have "k \ range abs" if "k \ \" for k :: int using that by induct simp ultimately show ?thesis by blast -qed +qed lemma Suc_nat_eq_nat_zadd1: "(0::int) <= z ==> Suc (nat z) = nat (1 + z)" apply (rule sym) diff -r a8a8eca85801 -r 3e491e34a62e src/HOL/Multivariate_Analysis/Cauchy_Integral_Thm.thy --- a/src/HOL/Multivariate_Analysis/Cauchy_Integral_Thm.thy Mon Sep 21 14:44:32 2015 +0200 +++ b/src/HOL/Multivariate_Analysis/Cauchy_Integral_Thm.thy Mon Sep 21 19:52:13 2015 +0100 @@ -4,58 +4,6 @@ imports Complex_Transcendental Weierstrass begin -(*FIXME migrate into libraries*) - -lemma inj_mult_left [simp]: "inj (op* x) \ x \ (0::'a::idom)" - by (metis injI mult_cancel_left the_inv_f_f zero_neq_one) - -lemma continuous_on_strong_cong: - "s = t \ (\x. x \ t =simp=> f x = g x) \ continuous_on s f \ continuous_on t g" - by (simp add: simp_implies_def) - -thm image_affinity_atLeastAtMost_div_diff -lemma image_affinity_atLeastAtMost_div: - fixes c :: "'a::linordered_field" - shows "((\x. x/m + c) ` {a..b}) = (if {a..b}={} then {} - else if 0 \ m then {a/m + c .. b/m + c} - else {b/m + c .. a/m + c})" - using image_affinity_atLeastAtMost [of "inverse m" c a b] - by (simp add: field_class.field_divide_inverse algebra_simps) - -thm continuous_on_closed_Un -lemma continuous_on_open_Un: - "open s \ open t \ continuous_on s f \ continuous_on t f \ continuous_on (s \ t) f" - using continuous_on_open_Union [of "{s,t}"] by auto - -thm continuous_on_eq (*REPLACE*) -lemma continuous_on_eq: - "\continuous_on s f; \x. x \ s \ f x = g x\ \ continuous_on s g" - unfolding continuous_on_def tendsto_def eventually_at_topological - by simp - -thm vector_derivative_add_at -lemma vector_derivative_mult_at [simp]: - fixes f g :: "real \ 'a :: real_normed_algebra" - shows "\f differentiable at a; g differentiable at a\ - \ vector_derivative (\x. f x * g x) (at a) = f a * vector_derivative g (at a) + vector_derivative f (at a) * g a" - by (simp add: vector_derivative_at has_vector_derivative_mult vector_derivative_works [symmetric]) - -lemma vector_derivative_scaleR_at [simp]: - "\f differentiable at a; g differentiable at a\ - \ vector_derivative (\x. f x *\<^sub>R g x) (at a) = f a *\<^sub>R vector_derivative g (at a) + vector_derivative f (at a) *\<^sub>R g a" -apply (rule vector_derivative_at) -apply (rule has_vector_derivative_scaleR) -apply (auto simp: vector_derivative_works has_vector_derivative_def has_field_derivative_def mult_commute_abs) -done - -thm Derivative.vector_diff_chain_at -lemma vector_derivative_chain_at: - assumes "f differentiable at x" "(g differentiable at (f x))" - shows "vector_derivative (g \ f) (at x) = - vector_derivative f (at x) *\<^sub>R vector_derivative g (at (f x))" -by (metis Derivative.vector_diff_chain_at vector_derivative_at vector_derivative_works assms) - - subsection \Piecewise differentiable functions\ definition piecewise_differentiable_on @@ -89,6 +37,9 @@ by (auto simp: piecewise_differentiable_on_def differentiable_imp_continuous_on differentiable_on_def intro: differentiable_within_subset) +lemma piecewise_differentiable_const [iff]: "(\x. z) piecewise_differentiable_on s" + by (simp add: differentiable_imp_piecewise_differentiable) + lemma piecewise_differentiable_compose: "\f piecewise_differentiable_on s; g piecewise_differentiable_on (f ` s); \x. finite (s \ f-`{x})\ @@ -466,7 +417,7 @@ by (auto simp: piecewise_C1_differentiable_on_def continuous_on_add) qed -lemma piecewise_C1_differentiable_C1_diff: +lemma piecewise_C1_differentiable_diff: "\f piecewise_C1_differentiable_on s; g piecewise_C1_differentiable_on s\ \ (\x. f x - g x) piecewise_C1_differentiable_on s" unfolding diff_conv_add_uminus diff -r a8a8eca85801 -r 3e491e34a62e src/HOL/Multivariate_Analysis/Derivative.thy --- a/src/HOL/Multivariate_Analysis/Derivative.thy Mon Sep 21 14:44:32 2015 +0200 +++ b/src/HOL/Multivariate_Analysis/Derivative.thy Mon Sep 21 19:52:13 2015 +0100 @@ -2277,6 +2277,20 @@ \ vector_derivative (\x. f x - g x) (at a) = vector_derivative f (at a) - vector_derivative g (at a)" by (simp add: vector_derivative_at has_vector_derivative_diff vector_derivative_works [symmetric]) +lemma vector_derivative_mult_at [simp]: + fixes f g :: "real \ 'a :: real_normed_algebra" + shows "\f differentiable at a; g differentiable at a\ + \ vector_derivative (\x. f x * g x) (at a) = f a * vector_derivative g (at a) + vector_derivative f (at a) * g a" + by (simp add: vector_derivative_at has_vector_derivative_mult vector_derivative_works [symmetric]) + +lemma vector_derivative_scaleR_at [simp]: + "\f differentiable at a; g differentiable at a\ + \ vector_derivative (\x. f x *\<^sub>R g x) (at a) = f a *\<^sub>R vector_derivative g (at a) + vector_derivative f (at a) *\<^sub>R g a" +apply (rule vector_derivative_at) +apply (rule has_vector_derivative_scaleR) +apply (auto simp: vector_derivative_works has_vector_derivative_def has_field_derivative_def mult_commute_abs) +done + lemma vector_derivative_within_closed_interval: assumes "a < b" and "x \ cbox a b" @@ -2356,4 +2370,10 @@ a \ x \ x \ b \ a vector_derivative f (at x within {a..b}) = f'" using has_vector_derivative_at_within vector_derivative_within_closed_interval by fastforce +lemma vector_derivative_chain_at: + assumes "f differentiable at x" "(g differentiable at (f x))" + shows "vector_derivative (g \ f) (at x) = + vector_derivative f (at x) *\<^sub>R vector_derivative g (at (f x))" +by (metis vector_diff_chain_at vector_derivative_at vector_derivative_works assms) + end diff -r a8a8eca85801 -r 3e491e34a62e src/HOL/Multivariate_Analysis/Integration.thy --- a/src/HOL/Multivariate_Analysis/Integration.thy Mon Sep 21 14:44:32 2015 +0200 +++ b/src/HOL/Multivariate_Analysis/Integration.thy Mon Sep 21 19:52:13 2015 +0100 @@ -496,6 +496,9 @@ lemma content_real: "a \ b \ content {a..b} = b - a" by (auto simp: interval_upperbound_def interval_lowerbound_def SUP_def INF_def content_def) +lemma abs_eq_content: "abs (y - x) = (if x\y then content {x .. y} else content {y..x})" + by (auto simp: content_real) + lemma content_singleton[simp]: "content {a} = 0" proof - have "content (cbox a a) = 0" @@ -6414,133 +6417,75 @@ apply (rule has_integral_const) done +lemma integral_has_vector_derivative_continuous_at: + fixes f :: "real \ 'a::banach" + assumes f: "f integrable_on {a..b}" + and x: "x \ {a..b}" + and fx: "continuous (at x within {a..b}) f" + shows "((\u. integral {a..u} f) has_vector_derivative f x) (at x within {a..b})" +proof - + let ?I = "\a b. integral {a..b} f" + { fix e::real + assume "e > 0" + obtain d where "d>0" and d: "\x'. \x' \ {a..b}; \x' - x\ < d\ \ norm(f x' - f x) \ e" + using `e>0` fx by (auto simp: continuous_within_eps_delta dist_norm less_imp_le) + have "norm (integral {a..y} f - integral {a..x} f - (y - x) *\<^sub>R f x) \ e * \y - x\" + if y: "y \ {a..b}" and yx: "\y - x\ < d" for y + proof (cases "y < x") + case False + have "f integrable_on {a..y}" + using f y by (simp add: integrable_subinterval_real) + then have Idiff: "?I a y - ?I a x = ?I x y" + using False x by (simp add: algebra_simps integral_combine) + have fux_int: "((\u. f u - f x) has_integral integral {x..y} f - (y - x) *\<^sub>R f x) {x..y}" + apply (rule has_integral_sub) + using x y apply (force intro: integrable_integral [OF integrable_subinterval_real [OF f]]) + using has_integral_const_real [of "f x" x y] False + apply (simp add: ) + done + show ?thesis + using False + apply (simp add: abs_eq_content del: content_real_if) + apply (rule has_integral_bound_real[where f="(\u. f u - f x)"]) + using yx False d x y `e>0` apply (auto simp add: Idiff fux_int) + done + next + case True + have "f integrable_on {a..x}" + using f x by (simp add: integrable_subinterval_real) + then have Idiff: "?I a x - ?I a y = ?I y x" + using True x y by (simp add: algebra_simps integral_combine) + have fux_int: "((\u. f u - f x) has_integral integral {y..x} f - (x - y) *\<^sub>R f x) {y..x}" + apply (rule has_integral_sub) + using x y apply (force intro: integrable_integral [OF integrable_subinterval_real [OF f]]) + using has_integral_const_real [of "f x" y x] True + apply (simp add: ) + done + have "norm (integral {a..x} f - integral {a..y} f - (x - y) *\<^sub>R f x) \ e * \y - x\" + using True + apply (simp add: abs_eq_content del: content_real_if) + apply (rule has_integral_bound_real[where f="(\u. f u - f x)"]) + using yx True d x y `e>0` apply (auto simp add: Idiff fux_int) + done + then show ?thesis + by (simp add: algebra_simps norm_minus_commute) + qed + then have "\d>0. \y\{a..b}. \y - x\ < d \ norm (integral {a..y} f - integral {a..x} f - (y - x) *\<^sub>R f x) \ e * \y - x\" + using `d>0` by blast + } + then show ?thesis + by (simp add: has_vector_derivative_def has_derivative_within_alt bounded_linear_scaleR_left) +qed + lemma integral_has_vector_derivative: fixes f :: "real \ 'a::banach" assumes "continuous_on {a .. b} f" and "x \ {a .. b}" shows "((\u. integral {a .. u} f) has_vector_derivative f(x)) (at x within {a .. b})" - unfolding has_vector_derivative_def has_derivative_within_alt - apply safe - apply (rule bounded_linear_scaleR_left) -proof - - fix e :: real - assume e: "e > 0" - note compact_uniformly_continuous[OF assms(1) compact_Icc,unfolded uniformly_continuous_on_def] - from this[rule_format,OF e] guess d by (elim conjE exE) note d=this[rule_format] - let ?I = "\a b. integral {a .. b} f" - show "\d>0. \y\{a .. b}. norm (y - x) < d \ - norm (?I a y - ?I a x - (y - x) *\<^sub>R f x) \ e * norm (y - x)" - proof (rule, rule, rule d, safe, goal_cases) - case prems: (1 y) - show ?case - proof (cases "y < x") - case False - have "f integrable_on {a .. y}" - apply (rule integrable_subinterval_real,rule integrable_continuous_real) - apply (rule assms) - unfolding not_less - using assms(2) prems - apply auto - done - then have *: "?I a y - ?I a x = ?I x y" - unfolding algebra_simps - apply (subst eq_commute) - apply (rule integral_combine) - using False - unfolding not_less - using assms(2) prems - apply auto - done - have **: "norm (y - x) = content {x .. y}" - using False by (auto simp: content_real) - show ?thesis - unfolding ** - apply (rule has_integral_bound_real[where f="(\u. f u - f x)"]) - unfolding * - defer - apply (rule has_integral_sub) - apply (rule integrable_integral) - apply (rule integrable_subinterval_real) - apply (rule integrable_continuous_real) - apply (rule assms)+ - proof - - show "{x .. y} \ {a .. b}" - using prems assms(2) by auto - have *: "y - x = norm (y - x)" - using False by auto - show "((\xa. f x) has_integral (y - x) *\<^sub>R f x) {x .. y}" - apply (subst *) - unfolding ** - by blast - show "\xa\{x .. y}. norm (f xa - f x) \ e" - apply safe - apply (rule less_imp_le) - apply (rule d(2)[unfolded dist_norm]) - using assms(2) - using prems - apply auto - done - qed (insert e, auto) - next - case True - have "f integrable_on cbox a x" - apply (rule integrable_subinterval,rule integrable_continuous) - unfolding box_real - apply (rule assms)+ - unfolding not_less - using assms(2) prems - apply auto - done - then have *: "?I a x - ?I a y = ?I y x" - unfolding algebra_simps - apply (subst eq_commute) - apply (rule integral_combine) - using True using assms(2) prems - apply auto - done - have **: "norm (y - x) = content {y .. x}" - apply (subst content_real) - using True - unfolding not_less - apply auto - done - have ***: "\fy fx c::'a. fx - fy - (y - x) *\<^sub>R c = -(fy - fx - (x - y) *\<^sub>R c)" - unfolding scaleR_left.diff by auto - show ?thesis - apply (subst ***) - unfolding norm_minus_cancel ** - apply (rule has_integral_bound_real[where f="(\u. f u - f x)"]) - unfolding * - unfolding o_def - defer - apply (rule has_integral_sub) - apply (subst minus_minus[symmetric]) - unfolding minus_minus - apply (rule integrable_integral) - apply (rule integrable_subinterval_real,rule integrable_continuous_real) - apply (rule assms)+ - proof - - show "{y .. x} \ {a .. b}" - using prems assms(2) by auto - have *: "x - y = norm (y - x)" - using True by auto - show "((\xa. f x) has_integral (x - y) *\<^sub>R f x) {y .. x}" - apply (subst *) - unfolding ** - apply blast - done - show "\xa\{y .. x}. norm (f xa - f x) \ e" - apply safe - apply (rule less_imp_le) - apply (rule d(2)[unfolded dist_norm]) - using assms(2) - using prems - apply auto - done - qed (insert e, auto) - qed - qed -qed +apply (rule integral_has_vector_derivative_continuous_at [OF integrable_continuous_real]) +using assms +apply (auto simp: continuous_on_eq_continuous_within) +done lemma antiderivative_continuous: fixes q b :: real diff -r a8a8eca85801 -r 3e491e34a62e src/HOL/Multivariate_Analysis/Path_Connected.thy --- a/src/HOL/Multivariate_Analysis/Path_Connected.thy Mon Sep 21 14:44:32 2015 +0200 +++ b/src/HOL/Multivariate_Analysis/Path_Connected.thy Mon Sep 21 19:52:13 2015 +0100 @@ -437,11 +437,8 @@ apply (auto simp: continuous_on_op_minus) done -lemma forall_01_trivial: "(\x\{0..1}. x \ 0 \ P x) \ P (0::real)" - by auto - -lemma forall_half1_trivial: "(\x\{1/2..1}. x * 2 \ 1 \ P x) \ P (1/2::real)" - by auto (metis add_divide_distrib mult_2_right real_sum_of_halves) +lemma half_bounded_equal: "1 \ x * 2 \ x * 2 \ 1 \ x = (1/2::real)" + by simp lemma continuous_on_joinpaths: assumes "continuous_on {0..1} g1" "continuous_on {0..1} g2" "pathfinish g1 = pathstart g2" @@ -451,17 +448,17 @@ by auto have gg: "g2 0 = g1 1" by (metis assms(3) pathfinish_def pathstart_def) - have 1: "continuous_on {0..1 / 2} (g1 +++ g2)" + have 1: "continuous_on {0..1/2} (g1 +++ g2)" apply (rule continuous_on_eq [of _ "g1 o (\x. 2*x)"]) - apply (simp add: joinpaths_def) - apply (rule continuous_intros | simp add: assms)+ + apply (rule continuous_intros | simp add: joinpaths_def assms)+ done - have 2: "continuous_on {1 / 2..1} (g1 +++ g2)" - apply (rule continuous_on_eq [of _ "g2 o (\x. 2*x-1)"]) - apply (simp add: joinpaths_def) - apply (rule continuous_intros | simp add: forall_half1_trivial gg)+ - apply (rule continuous_on_subset) - apply (rule assms, auto) + have "continuous_on {1/2..1} (g2 o (\x. 2*x-1))" + apply (rule continuous_on_subset [of "{1/2..1}"]) + apply (rule continuous_intros | simp add: image_affinity_atLeastAtMost_diff assms)+ + done + then have 2: "continuous_on {1/2..1} (g1 +++ g2)" + apply (rule continuous_on_eq [of "{1/2..1}" "g2 o (\x. 2*x-1)"]) + apply (rule assms continuous_intros | simp add: joinpaths_def mult.commute half_bounded_equal gg)+ done show ?thesis apply (subst *) @@ -800,7 +797,6 @@ apply (rule continuous_on_eq[of _ "g \ (\x. a + x)"]) prefer 3 apply (rule continuous_on_eq[of _ "g \ (\x. a - 1 + x)"]) - defer prefer 3 apply (rule continuous_intros)+ prefer 2 diff -r a8a8eca85801 -r 3e491e34a62e src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy --- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Mon Sep 21 14:44:32 2015 +0200 +++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Mon Sep 21 19:52:13 2015 +0100 @@ -2309,13 +2309,7 @@ subsection \More properties of closed balls\ -lemma closed_vimage: (* TODO: move to Topological_Spaces.thy *) - assumes "closed s" and "continuous_on UNIV f" - shows "closed (vimage f s)" - using assms unfolding continuous_on_closed_vimage [OF closed_UNIV] - by simp - -lemma closed_cball: "closed (cball x e)" +lemma closed_cball [iff]: "closed (cball x e)" proof - have "closed (dist x -` {..e})" by (intro closed_vimage closed_atMost continuous_intros) @@ -4663,7 +4657,7 @@ by (metis continuous_on_eq_continuous_at continuous_on_subset interiorE) lemma continuous_on_eq: - "(\x \ s. f x = g x) \ continuous_on s f \ continuous_on s g" + "\continuous_on s f; \x. x \ s \ f x = g x\ \ continuous_on s g" unfolding continuous_on_def tendsto_def eventually_at_topological by simp diff -r a8a8eca85801 -r 3e491e34a62e src/HOL/Real.thy --- a/src/HOL/Real.thy Mon Sep 21 14:44:32 2015 +0200 +++ b/src/HOL/Real.thy Mon Sep 21 19:52:13 2015 +0100 @@ -22,6 +22,13 @@ subsection \Preliminary lemmas\ +lemma inj_add_left [simp]: + fixes x :: "'a::cancel_semigroup_add" shows "inj (op+ x)" +by (meson add_left_imp_eq injI) + +lemma inj_mult_left [simp]: "inj (op* x) \ x \ (0::'a::idom)" + by (metis injI mult_cancel_left the_inv_f_f zero_neq_one) + lemma add_diff_add: fixes a b c d :: "'a::ab_group_add" shows "(a + c) - (b + d) = (a - b) + (c - d)" diff -r a8a8eca85801 -r 3e491e34a62e src/HOL/Set_Interval.thy --- a/src/HOL/Set_Interval.thy Mon Sep 21 14:44:32 2015 +0200 +++ b/src/HOL/Set_Interval.thy Mon Sep 21 19:52:13 2015 +0100 @@ -863,6 +863,14 @@ using image_affinity_atLeastAtMost [of m "-c" a b] by simp +lemma image_affinity_atLeastAtMost_div: + fixes c :: "'a::linordered_field" + shows "((\x. x/m + c) ` {a..b}) = (if {a..b}={} then {} + else if 0 \ m then {a/m + c .. b/m + c} + else {b/m + c .. a/m + c})" + using image_affinity_atLeastAtMost [of "inverse m" c a b] + by (simp add: field_class.field_divide_inverse algebra_simps) + lemma image_affinity_atLeastAtMost_div_diff: fixes c :: "'a::linordered_field" shows "((\x. x/m - c) ` {a..b}) = (if {a..b}={} then {} diff -r a8a8eca85801 -r 3e491e34a62e src/HOL/Topological_Spaces.thy --- a/src/HOL/Topological_Spaces.thy Mon Sep 21 14:44:32 2015 +0200 +++ b/src/HOL/Topological_Spaces.thy Mon Sep 21 19:52:13 2015 +0100 @@ -1392,6 +1392,10 @@ "(\s. s \ S \ open (A s)) \ (\s. s \ S \ continuous_on (A s) f) \ continuous_on (\s\S. A s) f" unfolding Union_image_eq[symmetric] by (rule continuous_on_open_Union) auto +lemma continuous_on_open_Un: + "open s \ open t \ continuous_on s f \ continuous_on t f \ continuous_on (s \ t) f" + using continuous_on_open_Union [of "{s,t}"] by auto + lemma continuous_on_closed_Un: "closed s \ closed t \ continuous_on s f \ continuous_on t f \ continuous_on (s \ t) f" by (auto simp add: continuous_on_closed_vimage closed_Un Int_Un_distrib)