| author | wenzelm | 
| Sat, 02 Apr 2016 23:14:08 +0200 | |
| changeset 62825 | e6e80a8bf624 | 
| parent 62533 | bc25f3916a99 | 
| child 62949 | f36a54da47a4 | 
| permissions | -rw-r--r-- | 
| 53781 | 1 | (* Title: HOL/Multivariate_Analysis/Derivative.thy | 
| 2 | Author: John Harrison | |
| 3 | Author: Robert Himmelmann, TU Muenchen (translation from HOL Light) | |
| 36350 | 4 | *) | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 5 | |
| 60420 | 6 | section \<open>Multivariate calculus in Euclidean space\<close> | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 7 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 8 | theory Derivative | 
| 62207 | 9 | imports Brouwer_Fixpoint Operator_Norm Uniform_Limit Bounded_Linear_Function | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 10 | begin | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 11 | |
| 61915 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61907diff
changeset | 12 | lemma onorm_inner_left: | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61907diff
changeset | 13 | assumes "bounded_linear r" | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61907diff
changeset | 14 | shows "onorm (\<lambda>x. r x \<bullet> f) \<le> onorm r * norm f" | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61907diff
changeset | 15 | proof (rule onorm_bound) | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61907diff
changeset | 16 | fix x | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61907diff
changeset | 17 | have "norm (r x \<bullet> f) \<le> norm (r x) * norm f" | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61907diff
changeset | 18 | by (simp add: Cauchy_Schwarz_ineq2) | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61907diff
changeset | 19 | also have "\<dots> \<le> onorm r * norm x * norm f" | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61907diff
changeset | 20 | by (intro mult_right_mono onorm assms norm_ge_zero) | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61907diff
changeset | 21 | finally show "norm (r x \<bullet> f) \<le> onorm r * norm f * norm x" | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61907diff
changeset | 22 | by (simp add: ac_simps) | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61907diff
changeset | 23 | qed (intro mult_nonneg_nonneg norm_ge_zero onorm_pos_le assms) | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61907diff
changeset | 24 | |
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61907diff
changeset | 25 | lemma onorm_inner_right: | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61907diff
changeset | 26 | assumes "bounded_linear r" | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61907diff
changeset | 27 | shows "onorm (\<lambda>x. f \<bullet> r x) \<le> norm f * onorm r" | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61907diff
changeset | 28 | apply (subst inner_commute) | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61907diff
changeset | 29 | apply (rule onorm_inner_left[OF assms, THEN order_trans]) | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61907diff
changeset | 30 | apply simp | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61907diff
changeset | 31 | done | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61907diff
changeset | 32 | |
| 56369 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 hoelzl parents: 
56332diff
changeset | 33 | declare has_derivative_bounded_linear[dest] | 
| 44137 | 34 | |
| 60420 | 35 | subsection \<open>Derivatives\<close> | 
| 51642 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 36 | |
| 60420 | 37 | subsubsection \<open>Combining theorems.\<close> | 
| 51642 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 38 | |
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
56151diff
changeset | 39 | lemmas has_derivative_id = has_derivative_ident | 
| 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
56151diff
changeset | 40 | lemmas has_derivative_neg = has_derivative_minus | 
| 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
56151diff
changeset | 41 | lemmas has_derivative_sub = has_derivative_diff | 
| 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
56151diff
changeset | 42 | lemmas scaleR_right_has_derivative = has_derivative_scaleR_right | 
| 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
56151diff
changeset | 43 | lemmas scaleR_left_has_derivative = has_derivative_scaleR_left | 
| 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
56151diff
changeset | 44 | lemmas inner_right_has_derivative = has_derivative_inner_right | 
| 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
56151diff
changeset | 45 | lemmas inner_left_has_derivative = has_derivative_inner_left | 
| 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
56151diff
changeset | 46 | lemmas mult_right_has_derivative = has_derivative_mult_right | 
| 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
56151diff
changeset | 47 | lemmas mult_left_has_derivative = has_derivative_mult_left | 
| 51642 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 48 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 49 | lemma has_derivative_add_const: | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 50 | "(f has_derivative f') net \<Longrightarrow> ((\<lambda>x. f x + c) has_derivative f') net" | 
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 51 | by (intro derivative_eq_intros) auto | 
| 51642 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 52 | |
| 53781 | 53 | |
| 60420 | 54 | subsection \<open>Derivative with composed bilinear function.\<close> | 
| 51642 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 55 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 56 | lemma has_derivative_bilinear_within: | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 57 | assumes "(f has_derivative f') (at x within s)" | 
| 53781 | 58 | and "(g has_derivative g') (at x within s)" | 
| 59 | and "bounded_bilinear h" | |
| 51642 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 60 | shows "((\<lambda>x. h (f x) (g x)) has_derivative (\<lambda>d. h (f x) (g' d) + h (f' d) (g x))) (at x within s)" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 61 | using bounded_bilinear.FDERIV[OF assms(3,1,2)] . | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 62 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 63 | lemma has_derivative_bilinear_at: | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 64 | assumes "(f has_derivative f') (at x)" | 
| 53781 | 65 | and "(g has_derivative g') (at x)" | 
| 66 | and "bounded_bilinear h" | |
| 51642 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 67 | shows "((\<lambda>x. h (f x) (g x)) has_derivative (\<lambda>d. h (f x) (g' d) + h (f' d) (g x))) (at x)" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 68 | using has_derivative_bilinear_within[of f f' x UNIV g g' h] assms by simp | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 69 | |
| 60420 | 70 | text \<open>These are the only cases we'll care about, probably.\<close> | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 71 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 72 | lemma has_derivative_within: "(f has_derivative f') (at x within s) \<longleftrightarrow> | 
| 61973 | 73 | bounded_linear f' \<and> ((\<lambda>y. (1 / norm(y - x)) *\<^sub>R (f y - (f x + f' (y - x)))) \<longlongrightarrow> 0) (at x within s)" | 
| 53781 | 74 | unfolding has_derivative_def Lim | 
| 57865 | 75 | by (auto simp add: netlimit_within field_simps) | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 76 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 77 | lemma has_derivative_at: "(f has_derivative f') (at x) \<longleftrightarrow> | 
| 61973 | 78 | bounded_linear f' \<and> ((\<lambda>y. (1 / (norm(y - x))) *\<^sub>R (f y - (f x + f' (y - x)))) \<longlongrightarrow> 0) (at x)" | 
| 53781 | 79 | using has_derivative_within [of f f' x UNIV] | 
| 80 | by simp | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 81 | |
| 60420 | 82 | text \<open>More explicit epsilon-delta forms.\<close> | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 83 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 84 | lemma has_derivative_within': | 
| 53781 | 85 | "(f has_derivative f')(at x within s) \<longleftrightarrow> | 
| 86 | bounded_linear f' \<and> | |
| 87 | (\<forall>e>0. \<exists>d>0. \<forall>x'\<in>s. 0 < norm (x' - x) \<and> norm (x' - x) < d \<longrightarrow> | |
| 88 | norm (f x' - f x - f'(x' - x)) / norm (x' - x) < e)" | |
| 36587 | 89 | unfolding has_derivative_within Lim_within dist_norm | 
| 53781 | 90 | unfolding diff_0_right | 
| 91 | by (simp add: diff_diff_eq) | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 92 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 93 | lemma has_derivative_at': | 
| 53781 | 94 | "(f has_derivative f') (at x) \<longleftrightarrow> bounded_linear f' \<and> | 
| 95 | (\<forall>e>0. \<exists>d>0. \<forall>x'. 0 < norm (x' - x) \<and> norm (x' - x) < d \<longrightarrow> | |
| 96 | norm (f x' - f x - f'(x' - x)) / norm (x' - x) < e)" | |
| 97 | using has_derivative_within' [of f f' x UNIV] | |
| 98 | by simp | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 99 | |
| 53781 | 100 | lemma has_derivative_at_within: | 
| 101 | "(f has_derivative f') (at x) \<Longrightarrow> (f has_derivative f') (at x within s)" | |
| 102 | unfolding has_derivative_within' has_derivative_at' | |
| 103 | by blast | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 104 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 105 | lemma has_derivative_within_open: | 
| 53781 | 106 | "a \<in> s \<Longrightarrow> open s \<Longrightarrow> | 
| 107 | (f has_derivative f') (at a within s) \<longleftrightarrow> (f has_derivative f') (at a)" | |
| 37730 | 108 | by (simp only: at_within_interior interior_open) | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 109 | |
| 43338 | 110 | lemma has_derivative_right: | 
| 53781 | 111 | fixes f :: "real \<Rightarrow> real" | 
| 112 | and y :: "real" | |
| 43338 | 113 |   shows "(f has_derivative (op * y)) (at x within ({x <..} \<inter> I)) \<longleftrightarrow>
 | 
| 61973 | 114 |     ((\<lambda>t. (f x - f t) / (x - t)) \<longlongrightarrow> y) (at x within ({x <..} \<inter> I))"
 | 
| 43338 | 115 | proof - | 
| 61973 | 116 |   have "((\<lambda>t. (f t - (f x + y * (t - x))) / \<bar>t - x\<bar>) \<longlongrightarrow> 0) (at x within ({x<..} \<inter> I)) \<longleftrightarrow>
 | 
| 117 |     ((\<lambda>t. (f t - f x) / (t - x) - y) \<longlongrightarrow> 0) (at x within ({x<..} \<inter> I))"
 | |
| 44282 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44140diff
changeset | 118 | by (intro Lim_cong_within) (auto simp add: diff_divide_distrib add_divide_distrib) | 
| 61973 | 119 |   also have "\<dots> \<longleftrightarrow> ((\<lambda>t. (f t - f x) / (t - x)) \<longlongrightarrow> y) (at x within ({x<..} \<inter> I))"
 | 
| 43338 | 120 | by (simp add: Lim_null[symmetric]) | 
| 61973 | 121 |   also have "\<dots> \<longleftrightarrow> ((\<lambda>t. (f x - f t) / (x - t)) \<longlongrightarrow> y) (at x within ({x<..} \<inter> I))"
 | 
| 44140 
2c10c35dd4be
remove several redundant and unused theorems about derivatives
 huffman parents: 
44137diff
changeset | 122 | by (intro Lim_cong_within) (simp_all add: field_simps) | 
| 43338 | 123 | finally show ?thesis | 
| 44282 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44140diff
changeset | 124 | by (simp add: bounded_linear_mult_right has_derivative_within) | 
| 43338 | 125 | qed | 
| 126 | ||
| 60420 | 127 | subsubsection \<open>Caratheodory characterization\<close> | 
| 55970 
6d123f0ae358
Some new proofs. Tidying up, esp to remove "apply rule".
 paulson <lp15@cam.ac.uk> parents: 
55665diff
changeset | 128 | |
| 
6d123f0ae358
Some new proofs. Tidying up, esp to remove "apply rule".
 paulson <lp15@cam.ac.uk> parents: 
55665diff
changeset | 129 | lemma DERIV_within_iff: | 
| 61973 | 130 | "(f has_field_derivative D) (at a within s) \<longleftrightarrow> ((\<lambda>z. (f z - f a) / (z - a)) \<longlongrightarrow> D) (at a within s)" | 
| 55970 
6d123f0ae358
Some new proofs. Tidying up, esp to remove "apply rule".
 paulson <lp15@cam.ac.uk> parents: 
55665diff
changeset | 131 | proof - | 
| 
6d123f0ae358
Some new proofs. Tidying up, esp to remove "apply rule".
 paulson <lp15@cam.ac.uk> parents: 
55665diff
changeset | 132 | have 1: "\<And>w y. ~(w = a) ==> y / (w - a) - D = (y - (w - a)*D)/(w - a)" | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57447diff
changeset | 133 | by (metis divide_diff_eq_iff eq_iff_diff_eq_0 mult.commute) | 
| 55970 
6d123f0ae358
Some new proofs. Tidying up, esp to remove "apply rule".
 paulson <lp15@cam.ac.uk> parents: 
55665diff
changeset | 134 | show ?thesis | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
56151diff
changeset | 135 | apply (simp add: has_field_derivative_def has_derivative_within bounded_linear_mult_right) | 
| 55970 
6d123f0ae358
Some new proofs. Tidying up, esp to remove "apply rule".
 paulson <lp15@cam.ac.uk> parents: 
55665diff
changeset | 136 | apply (simp add: LIM_zero_iff [where l = D, symmetric]) | 
| 
6d123f0ae358
Some new proofs. Tidying up, esp to remove "apply rule".
 paulson <lp15@cam.ac.uk> parents: 
55665diff
changeset | 137 | apply (simp add: Lim_within dist_norm) | 
| 
6d123f0ae358
Some new proofs. Tidying up, esp to remove "apply rule".
 paulson <lp15@cam.ac.uk> parents: 
55665diff
changeset | 138 | apply (simp add: nonzero_norm_divide [symmetric]) | 
| 59815 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59558diff
changeset | 139 | apply (simp add: 1 diff_diff_eq ac_simps) | 
| 55970 
6d123f0ae358
Some new proofs. Tidying up, esp to remove "apply rule".
 paulson <lp15@cam.ac.uk> parents: 
55665diff
changeset | 140 | done | 
| 
6d123f0ae358
Some new proofs. Tidying up, esp to remove "apply rule".
 paulson <lp15@cam.ac.uk> parents: 
55665diff
changeset | 141 | qed | 
| 
6d123f0ae358
Some new proofs. Tidying up, esp to remove "apply rule".
 paulson <lp15@cam.ac.uk> parents: 
55665diff
changeset | 142 | |
| 
6d123f0ae358
Some new proofs. Tidying up, esp to remove "apply rule".
 paulson <lp15@cam.ac.uk> parents: 
55665diff
changeset | 143 | lemma DERIV_caratheodory_within: | 
| 61104 
3c2d4636cebc
new lemmas about vector_derivative, complex numbers, paths, etc.
 paulson parents: 
61076diff
changeset | 144 | "(f has_field_derivative l) (at x within s) \<longleftrightarrow> | 
| 55970 
6d123f0ae358
Some new proofs. Tidying up, esp to remove "apply rule".
 paulson <lp15@cam.ac.uk> parents: 
55665diff
changeset | 145 | (\<exists>g. (\<forall>z. f z - f x = g z * (z - x)) \<and> continuous (at x within s) g \<and> g x = l)" | 
| 
6d123f0ae358
Some new proofs. Tidying up, esp to remove "apply rule".
 paulson <lp15@cam.ac.uk> parents: 
55665diff
changeset | 146 | (is "?lhs = ?rhs") | 
| 
6d123f0ae358
Some new proofs. Tidying up, esp to remove "apply rule".
 paulson <lp15@cam.ac.uk> parents: 
55665diff
changeset | 147 | proof | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
56151diff
changeset | 148 | assume ?lhs | 
| 55970 
6d123f0ae358
Some new proofs. Tidying up, esp to remove "apply rule".
 paulson <lp15@cam.ac.uk> parents: 
55665diff
changeset | 149 | show ?rhs | 
| 
6d123f0ae358
Some new proofs. Tidying up, esp to remove "apply rule".
 paulson <lp15@cam.ac.uk> parents: 
55665diff
changeset | 150 | proof (intro exI conjI) | 
| 
6d123f0ae358
Some new proofs. Tidying up, esp to remove "apply rule".
 paulson <lp15@cam.ac.uk> parents: 
55665diff
changeset | 151 | let ?g = "(%z. if z = x then l else (f z - f x) / (z-x))" | 
| 
6d123f0ae358
Some new proofs. Tidying up, esp to remove "apply rule".
 paulson <lp15@cam.ac.uk> parents: 
55665diff
changeset | 152 | show "\<forall>z. f z - f x = ?g z * (z-x)" by simp | 
| 60420 | 153 | show "continuous (at x within s) ?g" using \<open>?lhs\<close> | 
| 55970 
6d123f0ae358
Some new proofs. Tidying up, esp to remove "apply rule".
 paulson <lp15@cam.ac.uk> parents: 
55665diff
changeset | 154 | by (auto simp add: continuous_within DERIV_within_iff cong: Lim_cong_within) | 
| 
6d123f0ae358
Some new proofs. Tidying up, esp to remove "apply rule".
 paulson <lp15@cam.ac.uk> parents: 
55665diff
changeset | 155 | show "?g x = l" by simp | 
| 
6d123f0ae358
Some new proofs. Tidying up, esp to remove "apply rule".
 paulson <lp15@cam.ac.uk> parents: 
55665diff
changeset | 156 | qed | 
| 
6d123f0ae358
Some new proofs. Tidying up, esp to remove "apply rule".
 paulson <lp15@cam.ac.uk> parents: 
55665diff
changeset | 157 | next | 
| 
6d123f0ae358
Some new proofs. Tidying up, esp to remove "apply rule".
 paulson <lp15@cam.ac.uk> parents: 
55665diff
changeset | 158 | assume ?rhs | 
| 
6d123f0ae358
Some new proofs. Tidying up, esp to remove "apply rule".
 paulson <lp15@cam.ac.uk> parents: 
55665diff
changeset | 159 | then obtain g where | 
| 
6d123f0ae358
Some new proofs. Tidying up, esp to remove "apply rule".
 paulson <lp15@cam.ac.uk> parents: 
55665diff
changeset | 160 | "(\<forall>z. f z - f x = g z * (z-x))" and "continuous (at x within s) g" and "g x = l" by blast | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
56151diff
changeset | 161 | thus ?lhs | 
| 55970 
6d123f0ae358
Some new proofs. Tidying up, esp to remove "apply rule".
 paulson <lp15@cam.ac.uk> parents: 
55665diff
changeset | 162 | by (auto simp add: continuous_within DERIV_within_iff cong: Lim_cong_within) | 
| 
6d123f0ae358
Some new proofs. Tidying up, esp to remove "apply rule".
 paulson <lp15@cam.ac.uk> parents: 
55665diff
changeset | 163 | qed | 
| 
6d123f0ae358
Some new proofs. Tidying up, esp to remove "apply rule".
 paulson <lp15@cam.ac.uk> parents: 
55665diff
changeset | 164 | |
| 60420 | 165 | subsubsection \<open>Limit transformation for derivatives\<close> | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 166 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 167 | lemma has_derivative_transform_within: | 
| 62087 
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
 paulson parents: 
61975diff
changeset | 168 | assumes "(f has_derivative f') (at x within s)" | 
| 
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
 paulson parents: 
61975diff
changeset | 169 | and "0 < d" | 
| 53781 | 170 | and "x \<in> s" | 
| 62087 
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
 paulson parents: 
61975diff
changeset | 171 | and "\<And>x'. \<lbrakk>x' \<in> s; dist x' x < d\<rbrakk> \<Longrightarrow> f x' = g x'" | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 172 | shows "(g has_derivative f') (at x within s)" | 
| 55970 
6d123f0ae358
Some new proofs. Tidying up, esp to remove "apply rule".
 paulson <lp15@cam.ac.uk> parents: 
55665diff
changeset | 173 | using assms | 
| 62087 
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
 paulson parents: 
61975diff
changeset | 174 | unfolding has_derivative_within | 
| 
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
 paulson parents: 
61975diff
changeset | 175 | by (force simp add: intro: Lim_transform_within) | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 176 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 177 | lemma has_derivative_transform_within_open: | 
| 62087 
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
 paulson parents: 
61975diff
changeset | 178 | assumes "(f has_derivative f') (at x)" | 
| 
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
 paulson parents: 
61975diff
changeset | 179 | and "open s" | 
| 53781 | 180 | and "x \<in> s" | 
| 62087 
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
 paulson parents: 
61975diff
changeset | 181 | and "\<And>x. x\<in>s \<Longrightarrow> f x = g x" | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 182 | shows "(g has_derivative f') (at x)" | 
| 62087 
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
 paulson parents: 
61975diff
changeset | 183 | using assms unfolding has_derivative_at | 
| 
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
 paulson parents: 
61975diff
changeset | 184 | by (force simp add: intro: Lim_transform_within_open) | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 185 | |
| 60420 | 186 | subsection \<open>Differentiability\<close> | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 187 | |
| 53781 | 188 | definition | 
| 189 |   differentiable_on :: "('a::real_normed_vector \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> 'a set \<Rightarrow> bool"
 | |
| 56182 
528fae0816ea
update syntax of has_*derivative to infix 50; fixed proofs
 hoelzl parents: 
56181diff
changeset | 190 | (infix "differentiable'_on" 50) | 
| 53781 | 191 | where "f differentiable_on s \<longleftrightarrow> (\<forall>x\<in>s. f differentiable (at x within s))" | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 192 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 193 | lemma differentiableI: "(f has_derivative f') net \<Longrightarrow> f differentiable net" | 
| 53781 | 194 | unfolding differentiable_def | 
| 195 | by auto | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 196 | |
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62408diff
changeset | 197 | lemma differentiable_onD: "\<lbrakk>f differentiable_on S; x \<in> S\<rbrakk> \<Longrightarrow> f differentiable (at x within S)" | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62408diff
changeset | 198 | using differentiable_on_def by blast | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62408diff
changeset | 199 | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 200 | lemma differentiable_at_withinI: "f differentiable (at x) \<Longrightarrow> f differentiable (at x within s)" | 
| 53781 | 201 | unfolding differentiable_def | 
| 202 | using has_derivative_at_within | |
| 203 | by blast | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 204 | |
| 61104 
3c2d4636cebc
new lemmas about vector_derivative, complex numbers, paths, etc.
 paulson parents: 
61076diff
changeset | 205 | lemma differentiable_at_imp_differentiable_on: | 
| 
3c2d4636cebc
new lemmas about vector_derivative, complex numbers, paths, etc.
 paulson parents: 
61076diff
changeset | 206 | "(\<And>x. x \<in> s \<Longrightarrow> f differentiable at x) \<Longrightarrow> f differentiable_on s" | 
| 
3c2d4636cebc
new lemmas about vector_derivative, complex numbers, paths, etc.
 paulson parents: 
61076diff
changeset | 207 | by (metis differentiable_at_withinI differentiable_on_def) | 
| 
3c2d4636cebc
new lemmas about vector_derivative, complex numbers, paths, etc.
 paulson parents: 
61076diff
changeset | 208 | |
| 
3c2d4636cebc
new lemmas about vector_derivative, complex numbers, paths, etc.
 paulson parents: 
61076diff
changeset | 209 | corollary differentiable_iff_scaleR: | 
| 
3c2d4636cebc
new lemmas about vector_derivative, complex numbers, paths, etc.
 paulson parents: 
61076diff
changeset | 210 | fixes f :: "real \<Rightarrow> 'a::real_normed_vector" | 
| 
3c2d4636cebc
new lemmas about vector_derivative, complex numbers, paths, etc.
 paulson parents: 
61076diff
changeset | 211 | shows "f differentiable F \<longleftrightarrow> (\<exists>d. (f has_derivative (\<lambda>x. x *\<^sub>R d)) F)" | 
| 
3c2d4636cebc
new lemmas about vector_derivative, complex numbers, paths, etc.
 paulson parents: 
61076diff
changeset | 212 | by (auto simp: differentiable_def dest: has_derivative_linear linear_imp_scaleR) | 
| 
3c2d4636cebc
new lemmas about vector_derivative, complex numbers, paths, etc.
 paulson parents: 
61076diff
changeset | 213 | |
| 44123 | 214 | lemma differentiable_within_open: (* TODO: delete *) | 
| 53781 | 215 | assumes "a \<in> s" | 
| 216 | and "open s" | |
| 217 | shows "f differentiable (at a within s) \<longleftrightarrow> f differentiable (at a)" | |
| 218 | using assms | |
| 219 | by (simp only: at_within_interior interior_open) | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 220 | |
| 44123 | 221 | lemma differentiable_on_eq_differentiable_at: | 
| 53781 | 222 | "open s \<Longrightarrow> f differentiable_on s \<longleftrightarrow> (\<forall>x\<in>s. f differentiable at x)" | 
| 44123 | 223 | unfolding differentiable_on_def | 
| 51641 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51478diff
changeset | 224 | by (metis at_within_interior interior_open) | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 225 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 226 | lemma differentiable_transform_within: | 
| 62087 
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
 paulson parents: 
61975diff
changeset | 227 | assumes "f differentiable (at x within s)" | 
| 
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
 paulson parents: 
61975diff
changeset | 228 | and "0 < d" | 
| 53781 | 229 | and "x \<in> s" | 
| 62087 
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
 paulson parents: 
61975diff
changeset | 230 | and "\<And>x'. \<lbrakk>x'\<in>s; dist x' x < d\<rbrakk> \<Longrightarrow> f x' = g x'" | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 231 | shows "g differentiable (at x within s)" | 
| 62087 
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
 paulson parents: 
61975diff
changeset | 232 | using assms has_derivative_transform_within unfolding differentiable_def | 
| 
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
 paulson parents: 
61975diff
changeset | 233 | by blast | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 234 | |
| 53781 | 235 | |
| 60420 | 236 | subsection \<open>Frechet derivative and Jacobian matrix\<close> | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 237 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 238 | definition "frechet_derivative f net = (SOME f'. (f has_derivative f') net)" | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 239 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 240 | lemma frechet_derivative_works: | 
| 53781 | 241 | "f differentiable net \<longleftrightarrow> (f has_derivative (frechet_derivative f net)) net" | 
| 242 | unfolding frechet_derivative_def differentiable_def | |
| 243 | unfolding some_eq_ex[of "\<lambda> f' . (f has_derivative f') net"] .. | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 244 | |
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
56151diff
changeset | 245 | lemma linear_frechet_derivative: "f differentiable net \<Longrightarrow> linear (frechet_derivative f net)" | 
| 44123 | 246 | unfolding frechet_derivative_works has_derivative_def | 
| 56369 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 hoelzl parents: 
56332diff
changeset | 247 | by (auto intro: bounded_linear.linear) | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 248 | |
| 53781 | 249 | |
| 60420 | 250 | subsection \<open>Differentiability implies continuity\<close> | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 251 | |
| 44123 | 252 | lemma differentiable_imp_continuous_within: | 
| 51642 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 253 | "f differentiable (at x within s) \<Longrightarrow> continuous (at x within s) f" | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
56151diff
changeset | 254 | by (auto simp: differentiable_def intro: has_derivative_continuous) | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 255 | |
| 44123 | 256 | lemma differentiable_imp_continuous_on: | 
| 257 | "f differentiable_on s \<Longrightarrow> continuous_on s f" | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 258 | unfolding differentiable_on_def continuous_on_eq_continuous_within | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 259 | using differentiable_imp_continuous_within by blast | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 260 | |
| 44123 | 261 | lemma differentiable_on_subset: | 
| 262 | "f differentiable_on t \<Longrightarrow> s \<subseteq> t \<Longrightarrow> f differentiable_on s" | |
| 53781 | 263 | unfolding differentiable_on_def | 
| 264 | using differentiable_within_subset | |
| 265 | by blast | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 266 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 267 | lemma differentiable_on_empty: "f differentiable_on {}"
 | 
| 53781 | 268 | unfolding differentiable_on_def | 
| 269 | by auto | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 270 | |
| 60420 | 271 | text \<open>Results about neighborhoods filter.\<close> | 
| 56151 | 272 | |
| 273 | lemma eventually_nhds_metric_le: | |
| 274 | "eventually P (nhds a) = (\<exists>d>0. \<forall>x. dist x a \<le> d \<longrightarrow> P x)" | |
| 275 | unfolding eventually_nhds_metric by (safe, rule_tac x="d / 2" in exI, auto) | |
| 276 | ||
| 277 | lemma le_nhds: "F \<le> nhds a \<longleftrightarrow> (\<forall>S. open S \<and> a \<in> S \<longrightarrow> eventually (\<lambda>x. x \<in> S) F)" | |
| 61810 | 278 | unfolding le_filter_def eventually_nhds by (fast elim: eventually_mono) | 
| 56151 | 279 | |
| 280 | lemma le_nhds_metric: "F \<le> nhds a \<longleftrightarrow> (\<forall>e>0. eventually (\<lambda>x. dist x a < e) F)" | |
| 61810 | 281 | unfolding le_filter_def eventually_nhds_metric by (fast elim: eventually_mono) | 
| 56151 | 282 | |
| 283 | lemma le_nhds_metric_le: "F \<le> nhds a \<longleftrightarrow> (\<forall>e>0. eventually (\<lambda>x. dist x a \<le> e) F)" | |
| 61810 | 284 | unfolding le_filter_def eventually_nhds_metric_le by (fast elim: eventually_mono) | 
| 56151 | 285 | |
| 60420 | 286 | text \<open>Several results are easier using a "multiplied-out" variant. | 
| 287 | (I got this idea from Dieudonne's proof of the chain rule).\<close> | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 288 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 289 | lemma has_derivative_within_alt: | 
| 53781 | 290 | "(f has_derivative f') (at x within s) \<longleftrightarrow> bounded_linear f' \<and> | 
| 291 | (\<forall>e>0. \<exists>d>0. \<forall>y\<in>s. norm(y - x) < d \<longrightarrow> norm (f y - f x - f' (y - x)) \<le> e * norm (y - x))" | |
| 56151 | 292 | unfolding has_derivative_within filterlim_def le_nhds_metric_le eventually_filtermap | 
| 59815 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59558diff
changeset | 293 | eventually_at dist_norm diff_diff_eq | 
| 56369 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 hoelzl parents: 
56332diff
changeset | 294 | by (force simp add: linear_0 bounded_linear.linear pos_divide_le_eq) | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 295 | |
| 56320 | 296 | lemma has_derivative_within_alt2: | 
| 297 | "(f has_derivative f') (at x within s) \<longleftrightarrow> bounded_linear f' \<and> | |
| 298 | (\<forall>e>0. eventually (\<lambda>y. norm (f y - f x - f' (y - x)) \<le> e * norm (y - x)) (at x within s))" | |
| 299 | unfolding has_derivative_within filterlim_def le_nhds_metric_le eventually_filtermap | |
| 59815 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59558diff
changeset | 300 | eventually_at dist_norm diff_diff_eq | 
| 56369 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 hoelzl parents: 
56332diff
changeset | 301 | by (force simp add: linear_0 bounded_linear.linear pos_divide_le_eq) | 
| 56320 | 302 | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 303 | lemma has_derivative_at_alt: | 
| 53781 | 304 | "(f has_derivative f') (at x) \<longleftrightarrow> | 
| 305 | bounded_linear f' \<and> | |
| 306 | (\<forall>e>0. \<exists>d>0. \<forall>y. norm(y - x) < d \<longrightarrow> norm (f y - f x - f'(y - x)) \<le> e * norm (y - x))" | |
| 307 | using has_derivative_within_alt[where s=UNIV] | |
| 308 | by simp | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 309 | |
| 53781 | 310 | |
| 60420 | 311 | subsection \<open>The chain rule\<close> | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 312 | |
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 313 | lemma diff_chain_within[derivative_intros]: | 
| 44123 | 314 | assumes "(f has_derivative f') (at x within s)" | 
| 53781 | 315 | and "(g has_derivative g') (at (f x) within (f ` s))" | 
| 316 | shows "((g \<circ> f) has_derivative (g' \<circ> f'))(at x within s)" | |
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
56151diff
changeset | 317 | using has_derivative_in_compose[OF assms] | 
| 53781 | 318 | by (simp add: comp_def) | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 319 | |
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 320 | lemma diff_chain_at[derivative_intros]: | 
| 53781 | 321 | "(f has_derivative f') (at x) \<Longrightarrow> | 
| 322 | (g has_derivative g') (at (f x)) \<Longrightarrow> ((g \<circ> f) has_derivative (g' \<circ> f')) (at x)" | |
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
56151diff
changeset | 323 | using has_derivative_compose[of f f' x UNIV g g'] | 
| 53781 | 324 | by (simp add: comp_def) | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 325 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 326 | |
| 60420 | 327 | subsection \<open>Composition rules stated just for differentiability\<close> | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 328 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 329 | lemma differentiable_chain_at: | 
| 53781 | 330 | "f differentiable (at x) \<Longrightarrow> | 
| 331 | g differentiable (at (f x)) \<Longrightarrow> (g \<circ> f) differentiable (at x)" | |
| 332 | unfolding differentiable_def | |
| 333 | by (meson diff_chain_at) | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 334 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 335 | lemma differentiable_chain_within: | 
| 53781 | 336 | "f differentiable (at x within s) \<Longrightarrow> | 
| 337 | g differentiable (at(f x) within (f ` s)) \<Longrightarrow> (g \<circ> f) differentiable (at x within s)" | |
| 338 | unfolding differentiable_def | |
| 339 | by (meson diff_chain_within) | |
| 340 | ||
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 341 | |
| 60420 | 342 | subsection \<open>Uniqueness of derivative\<close> | 
| 37730 | 343 | |
| 56369 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 hoelzl parents: 
56332diff
changeset | 344 | |
| 60420 | 345 | text \<open> | 
| 37730 | 346 | The general result is a bit messy because we need approachability of the | 
| 347 | limit point from any direction. But OK for nontrivial intervals etc. | |
| 60420 | 348 | \<close> | 
| 51363 
d4d00c804645
changed has_derivative_intros into a named theorems collection
 hoelzl parents: 
50939diff
changeset | 349 | |
| 44123 | 350 | lemma frechet_derivative_unique_within: | 
| 351 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector" | |
| 352 | assumes "(f has_derivative f') (at x within s)" | |
| 53781 | 353 | and "(f has_derivative f'') (at x within s)" | 
| 61945 | 354 | and "\<forall>i\<in>Basis. \<forall>e>0. \<exists>d. 0 < \<bar>d\<bar> \<and> \<bar>d\<bar> < e \<and> (x + d *\<^sub>R i) \<in> s" | 
| 44123 | 355 | shows "f' = f''" | 
| 53781 | 356 | proof - | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 357 | note as = assms(1,2)[unfolded has_derivative_def] | 
| 44123 | 358 | then interpret f': bounded_linear f' by auto | 
| 359 | from as interpret f'': bounded_linear f'' by auto | |
| 360 | have "x islimpt s" unfolding islimpt_approachable | |
| 53781 | 361 | proof (rule, rule) | 
| 362 | fix e :: real | |
| 363 | assume "e > 0" | |
| 55665 | 364 | obtain d where "0 < \<bar>d\<bar>" and "\<bar>d\<bar> < e" and "x + d *\<^sub>R (SOME i. i \<in> Basis) \<in> s" | 
| 60420 | 365 | using assms(3) SOME_Basis \<open>e>0\<close> by blast | 
| 53781 | 366 | then show "\<exists>x'\<in>s. x' \<noteq> x \<and> dist x' x < e" | 
| 367 | apply (rule_tac x="x + d *\<^sub>R (SOME i. i \<in> Basis)" in bexI) | |
| 368 | unfolding dist_norm | |
| 369 | apply (auto simp: SOME_Basis nonzero_Basis) | |
| 370 | done | |
| 44123 | 371 | qed | 
| 53781 | 372 | then have *: "netlimit (at x within s) = x" | 
| 55970 
6d123f0ae358
Some new proofs. Tidying up, esp to remove "apply rule".
 paulson <lp15@cam.ac.uk> parents: 
55665diff
changeset | 373 | apply (auto intro!: netlimit_within) | 
| 
6d123f0ae358
Some new proofs. Tidying up, esp to remove "apply rule".
 paulson <lp15@cam.ac.uk> parents: 
55665diff
changeset | 374 | by (metis trivial_limit_within) | 
| 53781 | 375 | show ?thesis | 
| 376 | apply (rule linear_eq_stdbasis) | |
| 44123 | 377 | unfolding linear_conv_bounded_linear | 
| 53781 | 378 | apply (rule as(1,2)[THEN conjunct1])+ | 
| 379 | proof (rule, rule ccontr) | |
| 380 | fix i :: 'a | |
| 381 | assume i: "i \<in> Basis" | |
| 382 | def e \<equiv> "norm (f' i - f'' i)" | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50418diff
changeset | 383 | assume "f' i \<noteq> f'' i" | 
| 53781 | 384 | then have "e > 0" | 
| 385 | unfolding e_def by auto | |
| 55665 | 386 | obtain d where d: | 
| 387 | "0 < d" | |
| 388 | "(\<And>xa. xa\<in>s \<longrightarrow> 0 < dist xa x \<and> dist xa x < d \<longrightarrow> | |
| 389 | dist ((f xa - f x - f' (xa - x)) /\<^sub>R norm (xa - x) - | |
| 390 | (f xa - f x - f'' (xa - x)) /\<^sub>R norm (xa - x)) (0 - 0) < e)" | |
| 391 | using tendsto_diff [OF as(1,2)[THEN conjunct2]] | |
| 392 | unfolding * Lim_within | |
| 60420 | 393 | using \<open>e>0\<close> by blast | 
| 55665 | 394 | obtain c where c: "0 < \<bar>c\<bar>" "\<bar>c\<bar> < d \<and> x + c *\<^sub>R i \<in> s" | 
| 395 | using assms(3) i d(1) by blast | |
| 53781 | 396 | have *: "norm (- ((1 / \<bar>c\<bar>) *\<^sub>R f' (c *\<^sub>R i)) + (1 / \<bar>c\<bar>) *\<^sub>R f'' (c *\<^sub>R i)) = | 
| 61945 | 397 | norm ((1 / \<bar>c\<bar>) *\<^sub>R (- (f' (c *\<^sub>R i)) + f'' (c *\<^sub>R i)))" | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 398 | unfolding scaleR_right_distrib by auto | 
| 61945 | 399 | also have "\<dots> = norm ((1 / \<bar>c\<bar>) *\<^sub>R (c *\<^sub>R (- (f' i) + f'' i)))" | 
| 44123 | 400 | unfolding f'.scaleR f''.scaleR | 
| 53781 | 401 | unfolding scaleR_right_distrib scaleR_minus_right | 
| 402 | by auto | |
| 403 | also have "\<dots> = e" | |
| 404 | unfolding e_def | |
| 55665 | 405 | using c(1) | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50418diff
changeset | 406 | using norm_minus_cancel[of "f' i - f'' i"] | 
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
53799diff
changeset | 407 | by auto | 
| 53781 | 408 | finally show False | 
| 409 | using c | |
| 55665 | 410 | using d(2)[of "x + c *\<^sub>R i"] | 
| 44123 | 411 | unfolding dist_norm | 
| 412 | unfolding f'.scaleR f''.scaleR f'.add f''.add f'.diff f''.diff | |
| 413 | scaleR_scaleR scaleR_right_diff_distrib scaleR_right_distrib | |
| 53781 | 414 | using i | 
| 415 | by (auto simp: inverse_eq_divide) | |
| 44123 | 416 | qed | 
| 417 | qed | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 418 | |
| 37730 | 419 | lemma frechet_derivative_unique_at: | 
| 53781 | 420 | "(f has_derivative f') (at x) \<Longrightarrow> (f has_derivative f'') (at x) \<Longrightarrow> f' = f''" | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
56151diff
changeset | 421 | by (rule has_derivative_unique) | 
| 41829 | 422 | |
| 44123 | 423 | lemma frechet_derivative_unique_within_closed_interval: | 
| 56188 | 424 | fixes f::"'a::euclidean_space \<Rightarrow> 'b::real_normed_vector" | 
| 53781 | 425 | assumes "\<forall>i\<in>Basis. a\<bullet>i < b\<bullet>i" | 
| 56188 | 426 | and "x \<in> cbox a b" | 
| 427 | and "(f has_derivative f' ) (at x within cbox a b)" | |
| 428 | and "(f has_derivative f'') (at x within cbox a b)" | |
| 44123 | 429 | shows "f' = f''" | 
| 430 | apply(rule frechet_derivative_unique_within) | |
| 431 | apply(rule assms(3,4))+ | |
| 53781 | 432 | proof (rule, rule, rule) | 
| 433 | fix e :: real | |
| 434 | fix i :: 'a | |
| 435 | assume "e > 0" and i: "i \<in> Basis" | |
| 56188 | 436 | then show "\<exists>d. 0 < \<bar>d\<bar> \<and> \<bar>d\<bar> < e \<and> x + d *\<^sub>R i \<in> cbox a b" | 
| 53781 | 437 | proof (cases "x\<bullet>i = a\<bullet>i") | 
| 438 | case True | |
| 439 | then show ?thesis | |
| 440 | apply (rule_tac x="(min (b\<bullet>i - a\<bullet>i) e) / 2" in exI) | |
| 60420 | 441 | using assms(1)[THEN bspec[where x=i]] and \<open>e>0\<close> and assms(2) | 
| 56188 | 442 | unfolding mem_box | 
| 53781 | 443 | using i | 
| 444 | apply (auto simp add: field_simps inner_simps inner_Basis) | |
| 445 | done | |
| 446 | next | |
| 56188 | 447 | note * = assms(2)[unfolded mem_box, THEN bspec, OF i] | 
| 53781 | 448 | case False | 
| 449 | moreover have "a \<bullet> i < x \<bullet> i" | |
| 450 | using False * by auto | |
| 44123 | 451 |     moreover {
 | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50418diff
changeset | 452 | have "a \<bullet> i * 2 + min (x \<bullet> i - a \<bullet> i) e \<le> a\<bullet>i *2 + x\<bullet>i - a\<bullet>i" | 
| 44123 | 453 | by auto | 
| 53781 | 454 | also have "\<dots> = a\<bullet>i + x\<bullet>i" | 
| 455 | by auto | |
| 456 | also have "\<dots> \<le> 2 * (x\<bullet>i)" | |
| 457 | using * by auto | |
| 458 | finally have "a \<bullet> i * 2 + min (x \<bullet> i - a \<bullet> i) e \<le> x \<bullet> i * 2" | |
| 459 | by auto | |
| 44123 | 460 | } | 
| 53781 | 461 | moreover have "min (x \<bullet> i - a \<bullet> i) e \<ge> 0" | 
| 60420 | 462 | using * and \<open>e>0\<close> by auto | 
| 53781 | 463 | then have "x \<bullet> i * 2 \<le> b \<bullet> i * 2 + min (x \<bullet> i - a \<bullet> i) e" | 
| 464 | using * by auto | |
| 44123 | 465 | ultimately show ?thesis | 
| 53781 | 466 | apply (rule_tac x="- (min (x\<bullet>i - a\<bullet>i) e) / 2" in exI) | 
| 60420 | 467 | using assms(1)[THEN bspec, OF i] and \<open>e>0\<close> and assms(2) | 
| 56188 | 468 | unfolding mem_box | 
| 53781 | 469 | using i | 
| 470 | apply (auto simp add: field_simps inner_simps inner_Basis) | |
| 471 | done | |
| 44123 | 472 | qed | 
| 473 | qed | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 474 | |
| 44123 | 475 | lemma frechet_derivative_unique_within_open_interval: | 
| 56188 | 476 | fixes f::"'a::euclidean_space \<Rightarrow> 'b::real_normed_vector" | 
| 54775 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 immler parents: 
54230diff
changeset | 477 | assumes "x \<in> box a b" | 
| 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 immler parents: 
54230diff
changeset | 478 | and "(f has_derivative f' ) (at x within box a b)" | 
| 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 immler parents: 
54230diff
changeset | 479 | and "(f has_derivative f'') (at x within box a b)" | 
| 37650 | 480 | shows "f' = f''" | 
| 481 | proof - | |
| 54775 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 immler parents: 
54230diff
changeset | 482 | from assms(1) have *: "at x within box a b = at x" | 
| 56188 | 483 | by (metis at_within_interior interior_open open_box) | 
| 37650 | 484 | from assms(2,3) [unfolded *] show "f' = f''" | 
| 485 | by (rule frechet_derivative_unique_at) | |
| 486 | qed | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 487 | |
| 37730 | 488 | lemma frechet_derivative_at: | 
| 53781 | 489 | "(f has_derivative f') (at x) \<Longrightarrow> f' = frechet_derivative f (at x)" | 
| 490 | apply (rule frechet_derivative_unique_at[of f]) | |
| 491 | apply assumption | |
| 492 | unfolding frechet_derivative_works[symmetric] | |
| 493 | using differentiable_def | |
| 494 | apply auto | |
| 495 | done | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 496 | |
| 56188 | 497 | lemma frechet_derivative_within_cbox: | 
| 498 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector" | |
| 53781 | 499 | assumes "\<forall>i\<in>Basis. a\<bullet>i < b\<bullet>i" | 
| 56188 | 500 | and "x \<in> cbox a b" | 
| 501 | and "(f has_derivative f') (at x within cbox a b)" | |
| 502 | shows "frechet_derivative f (at x within cbox a b) = f'" | |
| 55970 
6d123f0ae358
Some new proofs. Tidying up, esp to remove "apply rule".
 paulson <lp15@cam.ac.uk> parents: 
55665diff
changeset | 503 | using assms | 
| 
6d123f0ae358
Some new proofs. Tidying up, esp to remove "apply rule".
 paulson <lp15@cam.ac.uk> parents: 
55665diff
changeset | 504 | by (metis Derivative.differentiableI frechet_derivative_unique_within_closed_interval frechet_derivative_works) | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 505 | |
| 53781 | 506 | |
| 60420 | 507 | subsection \<open>The traditional Rolle theorem in one dimension\<close> | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36844diff
changeset | 508 | |
| 60420 | 509 | text \<open>Derivatives of local minima and maxima are zero.\<close> | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 510 | |
| 56133 
304e37faf1ac
generalization of differential_zero_maxmin to class real_normed_vector
 huffman parents: 
56117diff
changeset | 511 | lemma has_derivative_local_min: | 
| 
304e37faf1ac
generalization of differential_zero_maxmin to class real_normed_vector
 huffman parents: 
56117diff
changeset | 512 | fixes f :: "'a::real_normed_vector \<Rightarrow> real" | 
| 
304e37faf1ac
generalization of differential_zero_maxmin to class real_normed_vector
 huffman parents: 
56117diff
changeset | 513 | assumes deriv: "(f has_derivative f') (at x)" | 
| 
304e37faf1ac
generalization of differential_zero_maxmin to class real_normed_vector
 huffman parents: 
56117diff
changeset | 514 | assumes min: "eventually (\<lambda>y. f x \<le> f y) (at x)" | 
| 
304e37faf1ac
generalization of differential_zero_maxmin to class real_normed_vector
 huffman parents: 
56117diff
changeset | 515 | shows "f' = (\<lambda>h. 0)" | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36844diff
changeset | 516 | proof | 
| 56133 
304e37faf1ac
generalization of differential_zero_maxmin to class real_normed_vector
 huffman parents: 
56117diff
changeset | 517 | fix h :: 'a | 
| 
304e37faf1ac
generalization of differential_zero_maxmin to class real_normed_vector
 huffman parents: 
56117diff
changeset | 518 | interpret f': bounded_linear f' | 
| 56182 
528fae0816ea
update syntax of has_*derivative to infix 50; fixed proofs
 hoelzl parents: 
56181diff
changeset | 519 | using deriv by (rule has_derivative_bounded_linear) | 
| 56133 
304e37faf1ac
generalization of differential_zero_maxmin to class real_normed_vector
 huffman parents: 
56117diff
changeset | 520 | show "f' h = 0" | 
| 
304e37faf1ac
generalization of differential_zero_maxmin to class real_normed_vector
 huffman parents: 
56117diff
changeset | 521 | proof (cases "h = 0") | 
| 
304e37faf1ac
generalization of differential_zero_maxmin to class real_normed_vector
 huffman parents: 
56117diff
changeset | 522 | assume "h \<noteq> 0" | 
| 
304e37faf1ac
generalization of differential_zero_maxmin to class real_normed_vector
 huffman parents: 
56117diff
changeset | 523 | from min obtain d where d1: "0 < d" and d2: "\<forall>y\<in>ball x d. f x \<le> f y" | 
| 
304e37faf1ac
generalization of differential_zero_maxmin to class real_normed_vector
 huffman parents: 
56117diff
changeset | 524 | unfolding eventually_at by (force simp: dist_commute) | 
| 
304e37faf1ac
generalization of differential_zero_maxmin to class real_normed_vector
 huffman parents: 
56117diff
changeset | 525 | have "FDERIV (\<lambda>r. x + r *\<^sub>R h) 0 :> (\<lambda>r. r *\<^sub>R h)" | 
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 526 | by (intro derivative_eq_intros) auto | 
| 56133 
304e37faf1ac
generalization of differential_zero_maxmin to class real_normed_vector
 huffman parents: 
56117diff
changeset | 527 | then have "FDERIV (\<lambda>r. f (x + r *\<^sub>R h)) 0 :> (\<lambda>k. f' (k *\<^sub>R h))" | 
| 56182 
528fae0816ea
update syntax of has_*derivative to infix 50; fixed proofs
 hoelzl parents: 
56181diff
changeset | 528 | by (rule has_derivative_compose, simp add: deriv) | 
| 56133 
304e37faf1ac
generalization of differential_zero_maxmin to class real_normed_vector
 huffman parents: 
56117diff
changeset | 529 | then have "DERIV (\<lambda>r. f (x + r *\<^sub>R h)) 0 :> f' h" | 
| 56182 
528fae0816ea
update syntax of has_*derivative to infix 50; fixed proofs
 hoelzl parents: 
56181diff
changeset | 530 | unfolding has_field_derivative_def by (simp add: f'.scaleR mult_commute_abs) | 
| 60420 | 531 | moreover have "0 < d / norm h" using d1 and \<open>h \<noteq> 0\<close> by simp | 
| 56133 
304e37faf1ac
generalization of differential_zero_maxmin to class real_normed_vector
 huffman parents: 
56117diff
changeset | 532 | moreover have "\<forall>y. \<bar>0 - y\<bar> < d / norm h \<longrightarrow> f (x + 0 *\<^sub>R h) \<le> f (x + y *\<^sub>R h)" | 
| 60420 | 533 | using \<open>h \<noteq> 0\<close> by (auto simp add: d2 dist_norm pos_less_divide_eq) | 
| 56133 
304e37faf1ac
generalization of differential_zero_maxmin to class real_normed_vector
 huffman parents: 
56117diff
changeset | 534 | ultimately show "f' h = 0" | 
| 
304e37faf1ac
generalization of differential_zero_maxmin to class real_normed_vector
 huffman parents: 
56117diff
changeset | 535 | by (rule DERIV_local_min) | 
| 
304e37faf1ac
generalization of differential_zero_maxmin to class real_normed_vector
 huffman parents: 
56117diff
changeset | 536 | qed (simp add: f'.zero) | 
| 
304e37faf1ac
generalization of differential_zero_maxmin to class real_normed_vector
 huffman parents: 
56117diff
changeset | 537 | qed | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36844diff
changeset | 538 | |
| 56133 
304e37faf1ac
generalization of differential_zero_maxmin to class real_normed_vector
 huffman parents: 
56117diff
changeset | 539 | lemma has_derivative_local_max: | 
| 
304e37faf1ac
generalization of differential_zero_maxmin to class real_normed_vector
 huffman parents: 
56117diff
changeset | 540 | fixes f :: "'a::real_normed_vector \<Rightarrow> real" | 
| 
304e37faf1ac
generalization of differential_zero_maxmin to class real_normed_vector
 huffman parents: 
56117diff
changeset | 541 | assumes "(f has_derivative f') (at x)" | 
| 
304e37faf1ac
generalization of differential_zero_maxmin to class real_normed_vector
 huffman parents: 
56117diff
changeset | 542 | assumes "eventually (\<lambda>y. f y \<le> f x) (at x)" | 
| 
304e37faf1ac
generalization of differential_zero_maxmin to class real_normed_vector
 huffman parents: 
56117diff
changeset | 543 | shows "f' = (\<lambda>h. 0)" | 
| 
304e37faf1ac
generalization of differential_zero_maxmin to class real_normed_vector
 huffman parents: 
56117diff
changeset | 544 | using has_derivative_local_min [of "\<lambda>x. - f x" "\<lambda>h. - f' h" "x"] | 
| 
304e37faf1ac
generalization of differential_zero_maxmin to class real_normed_vector
 huffman parents: 
56117diff
changeset | 545 | using assms unfolding fun_eq_iff by simp | 
| 
304e37faf1ac
generalization of differential_zero_maxmin to class real_normed_vector
 huffman parents: 
56117diff
changeset | 546 | |
| 
304e37faf1ac
generalization of differential_zero_maxmin to class real_normed_vector
 huffman parents: 
56117diff
changeset | 547 | lemma differential_zero_maxmin: | 
| 
304e37faf1ac
generalization of differential_zero_maxmin to class real_normed_vector
 huffman parents: 
56117diff
changeset | 548 | fixes f::"'a::real_normed_vector \<Rightarrow> real" | 
| 
304e37faf1ac
generalization of differential_zero_maxmin to class real_normed_vector
 huffman parents: 
56117diff
changeset | 549 | assumes "x \<in> s" | 
| 
304e37faf1ac
generalization of differential_zero_maxmin to class real_normed_vector
 huffman parents: 
56117diff
changeset | 550 | and "open s" | 
| 
304e37faf1ac
generalization of differential_zero_maxmin to class real_normed_vector
 huffman parents: 
56117diff
changeset | 551 | and deriv: "(f has_derivative f') (at x)" | 
| 
304e37faf1ac
generalization of differential_zero_maxmin to class real_normed_vector
 huffman parents: 
56117diff
changeset | 552 | and mono: "(\<forall>y\<in>s. f y \<le> f x) \<or> (\<forall>y\<in>s. f x \<le> f y)" | 
| 
304e37faf1ac
generalization of differential_zero_maxmin to class real_normed_vector
 huffman parents: 
56117diff
changeset | 553 | shows "f' = (\<lambda>v. 0)" | 
| 
304e37faf1ac
generalization of differential_zero_maxmin to class real_normed_vector
 huffman parents: 
56117diff
changeset | 554 | using mono | 
| 
304e37faf1ac
generalization of differential_zero_maxmin to class real_normed_vector
 huffman parents: 
56117diff
changeset | 555 | proof | 
| 
304e37faf1ac
generalization of differential_zero_maxmin to class real_normed_vector
 huffman parents: 
56117diff
changeset | 556 | assume "\<forall>y\<in>s. f y \<le> f x" | 
| 60420 | 557 | with \<open>x \<in> s\<close> and \<open>open s\<close> have "eventually (\<lambda>y. f y \<le> f x) (at x)" | 
| 56133 
304e37faf1ac
generalization of differential_zero_maxmin to class real_normed_vector
 huffman parents: 
56117diff
changeset | 558 | unfolding eventually_at_topological by auto | 
| 
304e37faf1ac
generalization of differential_zero_maxmin to class real_normed_vector
 huffman parents: 
56117diff
changeset | 559 | with deriv show ?thesis | 
| 
304e37faf1ac
generalization of differential_zero_maxmin to class real_normed_vector
 huffman parents: 
56117diff
changeset | 560 | by (rule has_derivative_local_max) | 
| 
304e37faf1ac
generalization of differential_zero_maxmin to class real_normed_vector
 huffman parents: 
56117diff
changeset | 561 | next | 
| 
304e37faf1ac
generalization of differential_zero_maxmin to class real_normed_vector
 huffman parents: 
56117diff
changeset | 562 | assume "\<forall>y\<in>s. f x \<le> f y" | 
| 60420 | 563 | with \<open>x \<in> s\<close> and \<open>open s\<close> have "eventually (\<lambda>y. f x \<le> f y) (at x)" | 
| 56133 
304e37faf1ac
generalization of differential_zero_maxmin to class real_normed_vector
 huffman parents: 
56117diff
changeset | 564 | unfolding eventually_at_topological by auto | 
| 
304e37faf1ac
generalization of differential_zero_maxmin to class real_normed_vector
 huffman parents: 
56117diff
changeset | 565 | with deriv show ?thesis | 
| 
304e37faf1ac
generalization of differential_zero_maxmin to class real_normed_vector
 huffman parents: 
56117diff
changeset | 566 | by (rule has_derivative_local_min) | 
| 
304e37faf1ac
generalization of differential_zero_maxmin to class real_normed_vector
 huffman parents: 
56117diff
changeset | 567 | qed | 
| 
304e37faf1ac
generalization of differential_zero_maxmin to class real_normed_vector
 huffman parents: 
56117diff
changeset | 568 | |
| 
304e37faf1ac
generalization of differential_zero_maxmin to class real_normed_vector
 huffman parents: 
56117diff
changeset | 569 | lemma differential_zero_maxmin_component: (* TODO: delete? *) | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36844diff
changeset | 570 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50418diff
changeset | 571 | assumes k: "k \<in> Basis" | 
| 53781 | 572 | and ball: "0 < e" "(\<forall>y \<in> ball x e. (f y)\<bullet>k \<le> (f x)\<bullet>k) \<or> (\<forall>y\<in>ball x e. (f x)\<bullet>k \<le> (f y)\<bullet>k)" | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36844diff
changeset | 573 | and diff: "f differentiable (at x)" | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50418diff
changeset | 574 | shows "(\<Sum>j\<in>Basis. (frechet_derivative f (at x) j \<bullet> k) *\<^sub>R j) = (0::'a)" (is "?D k = 0") | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36844diff
changeset | 575 | proof - | 
| 56133 
304e37faf1ac
generalization of differential_zero_maxmin to class real_normed_vector
 huffman parents: 
56117diff
changeset | 576 | let ?f' = "frechet_derivative f (at x)" | 
| 60420 | 577 | have "x \<in> ball x e" using \<open>0 < e\<close> by simp | 
| 56133 
304e37faf1ac
generalization of differential_zero_maxmin to class real_normed_vector
 huffman parents: 
56117diff
changeset | 578 | moreover have "open (ball x e)" by simp | 
| 
304e37faf1ac
generalization of differential_zero_maxmin to class real_normed_vector
 huffman parents: 
56117diff
changeset | 579 | moreover have "((\<lambda>x. f x \<bullet> k) has_derivative (\<lambda>h. ?f' h \<bullet> k)) (at x)" | 
| 
304e37faf1ac
generalization of differential_zero_maxmin to class real_normed_vector
 huffman parents: 
56117diff
changeset | 580 | using bounded_linear_inner_left diff[unfolded frechet_derivative_works] | 
| 56182 
528fae0816ea
update syntax of has_*derivative to infix 50; fixed proofs
 hoelzl parents: 
56181diff
changeset | 581 | by (rule bounded_linear.has_derivative) | 
| 56133 
304e37faf1ac
generalization of differential_zero_maxmin to class real_normed_vector
 huffman parents: 
56117diff
changeset | 582 | ultimately have "(\<lambda>h. frechet_derivative f (at x) h \<bullet> k) = (\<lambda>v. 0)" | 
| 
304e37faf1ac
generalization of differential_zero_maxmin to class real_normed_vector
 huffman parents: 
56117diff
changeset | 583 | using ball(2) by (rule differential_zero_maxmin) | 
| 
304e37faf1ac
generalization of differential_zero_maxmin to class real_normed_vector
 huffman parents: 
56117diff
changeset | 584 | then show ?thesis | 
| 
304e37faf1ac
generalization of differential_zero_maxmin to class real_normed_vector
 huffman parents: 
56117diff
changeset | 585 | unfolding fun_eq_iff by simp | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36844diff
changeset | 586 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 587 | |
| 51642 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 588 | lemma rolle: | 
| 53781 | 589 | fixes f :: "real \<Rightarrow> real" | 
| 590 | assumes "a < b" | |
| 591 | and "f a = f b" | |
| 56188 | 592 |     and "continuous_on {a .. b} f"
 | 
| 593 |     and "\<forall>x\<in>{a <..< b}. (f has_derivative f' x) (at x)"
 | |
| 594 |   shows "\<exists>x\<in>{a <..< b}. f' x = (\<lambda>v. 0)"
 | |
| 53781 | 595 | proof - | 
| 54775 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 immler parents: 
54230diff
changeset | 596 | have "\<exists>x\<in>box a b. (\<forall>y\<in>box a b. f x \<le> f y) \<or> (\<forall>y\<in>box a b. f y \<le> f x)" | 
| 53781 | 597 | proof - | 
| 598 |     have "(a + b) / 2 \<in> {a .. b}"
 | |
| 599 | using assms(1) by auto | |
| 56188 | 600 |     then have *: "{a .. b} \<noteq> {}"
 | 
| 53781 | 601 | by auto | 
| 55665 | 602 | obtain d where d: | 
| 56188 | 603 | "d \<in>cbox a b" | 
| 604 | "\<forall>y\<in>cbox a b. f y \<le> f d" | |
| 605 | using continuous_attains_sup[OF compact_Icc * assms(3)] by auto | |
| 55665 | 606 | obtain c where c: | 
| 56188 | 607 | "c \<in> cbox a b" | 
| 608 | "\<forall>y\<in>cbox a b. f c \<le> f y" | |
| 609 | using continuous_attains_inf[OF compact_Icc * assms(3)] by auto | |
| 44123 | 610 | show ?thesis | 
| 54775 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 immler parents: 
54230diff
changeset | 611 | proof (cases "d \<in> box a b \<or> c \<in> box a b") | 
| 53781 | 612 | case True | 
| 613 | then show ?thesis | |
| 56188 | 614 | by (metis c(2) d(2) box_subset_cbox subset_iff) | 
| 44123 | 615 | next | 
| 616 | def e \<equiv> "(a + b) /2" | |
| 53781 | 617 | case False | 
| 618 | then have "f d = f c" | |
| 56188 | 619 | using d c assms(2) by auto | 
| 53781 | 620 |       then have "\<And>x. x \<in> {a..b} \<Longrightarrow> f x = f d"
 | 
| 621 | using c d | |
| 55970 
6d123f0ae358
Some new proofs. Tidying up, esp to remove "apply rule".
 paulson <lp15@cam.ac.uk> parents: 
55665diff
changeset | 622 | by force | 
| 53781 | 623 | then show ?thesis | 
| 624 | apply (rule_tac x=e in bexI) | |
| 625 | unfolding e_def | |
| 626 | using assms(1) | |
| 56188 | 627 | apply auto | 
| 53781 | 628 | done | 
| 44123 | 629 | qed | 
| 630 | qed | |
| 56188 | 631 |   then obtain x where x: "x \<in> {a <..< b}" "(\<forall>y\<in>{a <..< b}. f x \<le> f y) \<or> (\<forall>y\<in>{a <..< b}. f y \<le> f x)"
 | 
| 632 | by auto | |
| 53781 | 633 | then have "f' x = (\<lambda>v. 0)" | 
| 54775 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 immler parents: 
54230diff
changeset | 634 | apply (rule_tac differential_zero_maxmin[of x "box a b" f "f' x"]) | 
| 55970 
6d123f0ae358
Some new proofs. Tidying up, esp to remove "apply rule".
 paulson <lp15@cam.ac.uk> parents: 
55665diff
changeset | 635 | using assms | 
| 53781 | 636 | apply auto | 
| 637 | done | |
| 638 | then show ?thesis | |
| 55970 
6d123f0ae358
Some new proofs. Tidying up, esp to remove "apply rule".
 paulson <lp15@cam.ac.uk> parents: 
55665diff
changeset | 639 | by (metis x(1)) | 
| 44123 | 640 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 641 | |
| 53781 | 642 | |
| 60420 | 643 | subsection \<open>One-dimensional mean value theorem\<close> | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 644 | |
| 53781 | 645 | lemma mvt: | 
| 646 | fixes f :: "real \<Rightarrow> real" | |
| 647 | assumes "a < b" | |
| 648 |     and "continuous_on {a..b} f"
 | |
| 44123 | 649 |   assumes "\<forall>x\<in>{a<..<b}. (f has_derivative (f' x)) (at x)"
 | 
| 53781 | 650 |   shows "\<exists>x\<in>{a<..<b}. f b - f a = (f' x) (b - a)"
 | 
| 651 | proof - | |
| 56188 | 652 |   have "\<exists>x\<in>{a <..< b}. (\<lambda>xa. f' x xa - (f b - f a) / (b - a) * xa) = (\<lambda>v. 0)"
 | 
| 51478 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 hoelzl parents: 
51363diff
changeset | 653 | proof (intro rolle[OF assms(1), of "\<lambda>x. f x - (f b - f a) / (b - a) * x"] ballI) | 
| 53781 | 654 | fix x | 
| 56188 | 655 |     assume x: "x \<in> {a <..< b}"
 | 
| 53781 | 656 | show "((\<lambda>x. f x - (f b - f a) / (b - a) * x) has_derivative | 
| 657 | (\<lambda>xa. f' x xa - (f b - f a) / (b - a) * xa)) (at x)" | |
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 658 | by (intro derivative_intros assms(3)[rule_format,OF x]) | 
| 56371 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 hoelzl parents: 
56370diff
changeset | 659 | qed (insert assms(1,2), auto intro!: continuous_intros simp: field_simps) | 
| 55665 | 660 | then obtain x where | 
| 56188 | 661 |     "x \<in> {a <..< b}"
 | 
| 55665 | 662 | "(\<lambda>xa. f' x xa - (f b - f a) / (b - a) * xa) = (\<lambda>v. 0)" .. | 
| 53781 | 663 | then show ?thesis | 
| 61762 
d50b993b4fb9
Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
 paulson <lp15@cam.ac.uk> parents: 
61649diff
changeset | 664 | by (metis (hide_lams) assms(1) diff_gt_0_iff_gt eq_iff_diff_eq_0 | 
| 59558 
5d9f0ace9af0
got rid of linordered_field_class.sign_simps(41) !
 paulson <lp15@cam.ac.uk> parents: 
58877diff
changeset | 665 | zero_less_mult_iff nonzero_mult_divide_cancel_right not_real_square_gt_zero | 
| 56188 | 666 | times_divide_eq_left) | 
| 44123 | 667 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 668 | |
| 44123 | 669 | lemma mvt_simple: | 
| 53781 | 670 | fixes f :: "real \<Rightarrow> real" | 
| 671 | assumes "a < b" | |
| 672 |     and "\<forall>x\<in>{a..b}. (f has_derivative f' x) (at x within {a..b})"
 | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 673 |   shows "\<exists>x\<in>{a<..<b}. f b - f a = f' x (b - a)"
 | 
| 56264 | 674 | proof (rule mvt) | 
| 675 |   have "f differentiable_on {a..b}"
 | |
| 676 | using assms(2) unfolding differentiable_on_def differentiable_def by fast | |
| 677 |   then show "continuous_on {a..b} f"
 | |
| 678 | by (rule differentiable_imp_continuous_on) | |
| 679 |   show "\<forall>x\<in>{a<..<b}. (f has_derivative f' x) (at x)"
 | |
| 680 | proof | |
| 681 | fix x | |
| 682 |     assume x: "x \<in> {a <..< b}"
 | |
| 683 | show "(f has_derivative f' x) (at x)" | |
| 684 | unfolding at_within_open[OF x open_greaterThanLessThan,symmetric] | |
| 685 | apply (rule has_derivative_within_subset) | |
| 686 | apply (rule assms(2)[rule_format]) | |
| 687 | using x | |
| 688 | apply auto | |
| 689 | done | |
| 690 | qed | |
| 691 | qed (rule assms(1)) | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 692 | |
| 44123 | 693 | lemma mvt_very_simple: | 
| 53781 | 694 | fixes f :: "real \<Rightarrow> real" | 
| 695 | assumes "a \<le> b" | |
| 56188 | 696 |     and "\<forall>x\<in>{a .. b}. (f has_derivative f' x) (at x within {a .. b})"
 | 
| 697 |   shows "\<exists>x\<in>{a .. b}. f b - f a = f' x (b - a)"
 | |
| 44123 | 698 | proof (cases "a = b") | 
| 53781 | 699 | interpret bounded_linear "f' b" | 
| 700 | using assms(2) assms(1) by auto | |
| 701 | case True | |
| 702 | then show ?thesis | |
| 703 | apply (rule_tac x=a in bexI) | |
| 704 | using assms(2)[THEN bspec[where x=a]] | |
| 705 | unfolding has_derivative_def | |
| 706 | unfolding True | |
| 707 | using zero | |
| 708 | apply auto | |
| 709 | done | |
| 710 | next | |
| 711 | case False | |
| 712 | then show ?thesis | |
| 713 | using mvt_simple[OF _ assms(2)] | |
| 714 | using assms(1) | |
| 715 | by auto | |
| 44123 | 716 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 717 | |
| 60420 | 718 | text \<open>A nice generalization (see Havin's proof of 5.19 from Rudin's book).\<close> | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 719 | |
| 44123 | 720 | lemma mvt_general: | 
| 56223 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
56217diff
changeset | 721 | fixes f :: "real \<Rightarrow> 'a::real_inner" | 
| 53781 | 722 | assumes "a < b" | 
| 56188 | 723 |     and "continuous_on {a .. b} f"
 | 
| 53781 | 724 |     and "\<forall>x\<in>{a<..<b}. (f has_derivative f'(x)) (at x)"
 | 
| 725 |   shows "\<exists>x\<in>{a<..<b}. norm (f b - f a) \<le> norm (f' x (b - a))"
 | |
| 726 | proof - | |
| 56264 | 727 |   have "\<exists>x\<in>{a<..<b}. (f b - f a) \<bullet> f b - (f b - f a) \<bullet> f a = (f b - f a) \<bullet> f' x (b - a)"
 | 
| 53781 | 728 | apply (rule mvt) | 
| 729 | apply (rule assms(1)) | |
| 56371 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 hoelzl parents: 
56370diff
changeset | 730 | apply (intro continuous_intros assms(2)) | 
| 53781 | 731 | using assms(3) | 
| 56264 | 732 | apply (fast intro: has_derivative_inner_right) | 
| 53781 | 733 | done | 
| 55665 | 734 | then obtain x where x: | 
| 735 |     "x \<in> {a<..<b}"
 | |
| 56264 | 736 | "(f b - f a) \<bullet> f b - (f b - f a) \<bullet> f a = (f b - f a) \<bullet> f' x (b - a)" .. | 
| 53781 | 737 | show ?thesis | 
| 738 | proof (cases "f a = f b") | |
| 36844 | 739 | case False | 
| 53077 | 740 | have "norm (f b - f a) * norm (f b - f a) = (norm (f b - f a))\<^sup>2" | 
| 44123 | 741 | by (simp add: power2_eq_square) | 
| 53781 | 742 | also have "\<dots> = (f b - f a) \<bullet> (f b - f a)" | 
| 743 | unfolding power2_norm_eq_inner .. | |
| 44123 | 744 | also have "\<dots> = (f b - f a) \<bullet> f' x (b - a)" | 
| 56264 | 745 | using x(2) by (simp only: inner_diff_right) | 
| 44123 | 746 | also have "\<dots> \<le> norm (f b - f a) * norm (f' x (b - a))" | 
| 747 | by (rule norm_cauchy_schwarz) | |
| 53781 | 748 | finally show ?thesis | 
| 749 | using False x(1) | |
| 56217 
dc429a5b13c4
Some rationalisation of basic lemmas
 paulson <lp15@cam.ac.uk> parents: 
56196diff
changeset | 750 | by (auto simp add: mult_left_cancel) | 
| 44123 | 751 | next | 
| 53781 | 752 | case True | 
| 753 | then show ?thesis | |
| 754 | using assms(1) | |
| 755 | apply (rule_tac x="(a + b) /2" in bexI) | |
| 756 | apply auto | |
| 757 | done | |
| 44123 | 758 | qed | 
| 759 | qed | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 760 | |
| 60178 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 761 | |
| 60420 | 762 | subsection \<open>More general bound theorems\<close> | 
| 60178 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 763 | |
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 764 | lemma differentiable_bound_general: | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 765 | fixes f :: "real \<Rightarrow> 'a::real_normed_vector" | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 766 | assumes "a < b" | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 767 |     and f_cont: "continuous_on {a .. b} f"
 | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 768 |     and phi_cont: "continuous_on {a .. b} \<phi>"
 | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 769 | and f': "\<And>x. a < x \<Longrightarrow> x < b \<Longrightarrow> (f has_vector_derivative f' x) (at x)" | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 770 | and phi': "\<And>x. a < x \<Longrightarrow> x < b \<Longrightarrow> (\<phi> has_vector_derivative \<phi>' x) (at x)" | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 771 | and bnd: "\<And>x. a < x \<Longrightarrow> x < b \<Longrightarrow> norm (f' x) \<le> \<phi>' x" | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 772 | shows "norm (f b - f a) \<le> \<phi> b - \<phi> a" | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 773 | proof - | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 774 |   {
 | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 775 | fix x assume x: "a < x" "x < b" | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 776 | have "0 \<le> norm (f' x)" by simp | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 777 | also have "\<dots> \<le> \<phi>' x" using x by (auto intro!: bnd) | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 778 | finally have "0 \<le> \<phi>' x" . | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 779 | } note phi'_nonneg = this | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 780 | note f_tendsto = assms(2)[simplified continuous_on_def, rule_format] | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 781 | note phi_tendsto = assms(3)[simplified continuous_on_def, rule_format] | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 782 |   {
 | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 783 | fix e::real assume "e > 0" | 
| 60420 | 784 | def e2 \<equiv> "e / 2" with \<open>e > 0\<close> have "e2 > 0" by simp | 
| 60178 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 785 | let ?le = "\<lambda>x1. norm (f x1 - f a) \<le> \<phi> x1 - \<phi> a + e * (x1 - a) + e" | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 786 |     def A \<equiv> "{x2. a \<le> x2 \<and> x2 \<le> b \<and> (\<forall>x1\<in>{a ..< x2}. ?le x1)}"
 | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 787 |     have A_subset: "A \<subseteq> {a .. b}" by (auto simp: A_def)
 | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 788 |     {
 | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 789 | fix x2 | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 790 |       assume a: "a \<le> x2" "x2 \<le> b" and le: "\<forall>x1\<in>{a..<x2}. ?le x1"
 | 
| 60420 | 791 | have "?le x2" using \<open>e > 0\<close> | 
| 60178 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 792 | proof cases | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 793 | assume "x2 \<noteq> a" with a have "a < x2" by simp | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 794 |         have "at x2 within {a <..<x2}\<noteq> bot"
 | 
| 60420 | 795 | using \<open>a < x2\<close> | 
| 60178 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 796 | by (auto simp: trivial_limit_within islimpt_in_closure) | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 797 | moreover | 
| 61973 | 798 |         have "((\<lambda>x1. (\<phi> x1 - \<phi> a) + e * (x1 - a) + e) \<longlongrightarrow> (\<phi> x2 - \<phi> a) + e * (x2 - a) + e) (at x2 within {a <..<x2})"
 | 
| 799 |           "((\<lambda>x1. norm (f x1 - f a)) \<longlongrightarrow> norm (f x2 - f a)) (at x2 within {a <..<x2})"
 | |
| 60178 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 800 | using a | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 801 | by (auto intro!: tendsto_eq_intros f_tendsto phi_tendsto | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 802 |             intro: tendsto_within_subset[where S="{a .. b}"])
 | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 803 | moreover | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 804 |         have "eventually (\<lambda>x. x > a) (at x2 within {a <..<x2})"
 | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 805 | by (auto simp: eventually_at_filter) | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 806 |         hence "eventually ?le (at x2 within {a <..<x2})"
 | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 807 | unfolding eventually_at_filter | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 808 | by eventually_elim (insert le, auto) | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 809 | ultimately | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 810 | show ?thesis | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 811 | by (rule tendsto_le) | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 812 | qed simp | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 813 | } note le_cont = this | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 814 | have "a \<in> A" | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 815 | using assms by (auto simp: A_def) | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 816 |     hence [simp]: "A \<noteq> {}" by auto
 | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 817 |     have A_ivl: "\<And>x1 x2. x2 \<in> A \<Longrightarrow> x1 \<in> {a ..x2} \<Longrightarrow> x1 \<in> A"
 | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 818 | by (simp add: A_def) | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 819 | have [simp]: "bdd_above A" by (auto simp: A_def) | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 820 | def y \<equiv> "Sup A" | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 821 | have "y \<le> b" | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 822 | unfolding y_def | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 823 | by (simp add: cSup_le_iff) (simp add: A_def) | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 824 | have leI: "\<And>x x1. a \<le> x1 \<Longrightarrow> x \<in> A \<Longrightarrow> x1 < x \<Longrightarrow> ?le x1" | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 825 | by (auto simp: A_def intro!: le_cont) | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 826 |     have y_all_le: "\<forall>x1\<in>{a..<y}. ?le x1"
 | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 827 | by (auto simp: y_def less_cSup_iff leI) | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 828 | have "a \<le> y" | 
| 60420 | 829 | by (metis \<open>a \<in> A\<close> \<open>bdd_above A\<close> cSup_upper y_def) | 
| 60178 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 830 | have "y \<in> A" | 
| 60420 | 831 | using y_all_le \<open>a \<le> y\<close> \<open>y \<le> b\<close> | 
| 60178 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 832 | by (auto simp: A_def) | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 833 |     hence "A = {a .. y}"
 | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 834 | using A_subset | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 835 | by (auto simp: subset_iff y_def cSup_upper intro: A_ivl) | 
| 60420 | 836 | from le_cont[OF \<open>a \<le> y\<close> \<open>y \<le> b\<close> y_all_le] have le_y: "?le y" . | 
| 60178 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 837 |     {
 | 
| 60420 | 838 | assume "a \<noteq> y" with \<open>a \<le> y\<close> have "a < y" by simp | 
| 60178 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 839 | have "y = b" | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 840 | proof (rule ccontr) | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 841 | assume "y \<noteq> b" | 
| 60420 | 842 | hence "y < b" using \<open>y \<le> b\<close> by simp | 
| 60178 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 843 |         let ?F = "at y within {y..<b}"
 | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 844 | from f' phi' | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 845 | have "(f has_vector_derivative f' y) ?F" | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 846 | and "(\<phi> has_vector_derivative \<phi>' y) ?F" | 
| 60420 | 847 | using \<open>a < y\<close> \<open>y < b\<close> | 
| 60178 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 848 |           by (auto simp add: at_within_open[of _ "{a<..<b}"] has_vector_derivative_def
 | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 849 |             intro!: has_derivative_subset[where s="{a<..<b}" and t="{y..<b}"])
 | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 850 | hence "\<forall>\<^sub>F x1 in ?F. norm (f x1 - f y - (x1 - y) *\<^sub>R f' y) \<le> e2 * \<bar>x1 - y\<bar>" | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 851 | "\<forall>\<^sub>F x1 in ?F. norm (\<phi> x1 - \<phi> y - (x1 - y) *\<^sub>R \<phi>' y) \<le> e2 * \<bar>x1 - y\<bar>" | 
| 60420 | 852 | using \<open>e2 > 0\<close> | 
| 60178 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 853 | by (auto simp: has_derivative_within_alt2 has_vector_derivative_def) | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 854 | moreover | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 855 | have "\<forall>\<^sub>F x1 in ?F. y \<le> x1" "\<forall>\<^sub>F x1 in ?F. x1 < b" | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 856 | by (auto simp: eventually_at_filter) | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 857 | ultimately | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 858 | have "\<forall>\<^sub>F x1 in ?F. norm (f x1 - f y) \<le> (\<phi> x1 - \<phi> y) + e * \<bar>x1 - y\<bar>" | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 859 | (is "\<forall>\<^sub>F x1 in ?F. ?le' x1") | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 860 | proof eventually_elim | 
| 60589 | 861 | case (elim x1) | 
| 60178 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 862 | from norm_triangle_ineq2[THEN order_trans, OF elim(1)] | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 863 | have "norm (f x1 - f y) \<le> norm (f' y) * \<bar>x1 - y\<bar> + e2 * \<bar>x1 - y\<bar>" | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 864 | by (simp add: ac_simps) | 
| 60420 | 865 | also have "norm (f' y) \<le> \<phi>' y" using bnd \<open>a < y\<close> \<open>y < b\<close> by simp | 
| 60178 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 866 | also | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 867 | from elim have "\<phi>' y * \<bar>x1 - y\<bar> \<le> \<phi> x1 - \<phi> y + e2 * \<bar>x1 - y\<bar>" | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 868 | by (simp add: ac_simps) | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 869 | finally | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 870 | have "norm (f x1 - f y) \<le> \<phi> x1 - \<phi> y + e2 * \<bar>x1 - y\<bar> + e2 * \<bar>x1 - y\<bar>" | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 871 | by (auto simp: mult_right_mono) | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 872 | thus ?case by (simp add: e2_def) | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 873 | qed | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 874 | moreover have "?le' y" by simp | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 875 | ultimately obtain S | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 876 |         where S: "open S" "y \<in> S" "\<And>x. x\<in>S \<Longrightarrow> x \<in> {y..<b} \<Longrightarrow> ?le' x"
 | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 877 | unfolding eventually_at_topological | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 878 | by metis | 
| 60420 | 879 | from \<open>open S\<close> obtain d where d: "\<And>x. dist x y < d \<Longrightarrow> x \<in> S" "d > 0" | 
| 62101 | 880 | by (force simp: dist_commute open_dist ball_def dest!: bspec[OF _ \<open>y \<in> S\<close>]) | 
| 60178 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 881 | def d' \<equiv> "min ((y + b)/2) (y + (d/2))" | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 882 | have "d' \<in> A" | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 883 | unfolding A_def | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 884 | proof safe | 
| 60420 | 885 | show "a \<le> d'" using \<open>a < y\<close> \<open>0 < d\<close> \<open>y < b\<close> by (simp add: d'_def) | 
| 886 | show "d' \<le> b" using \<open>y < b\<close> by (simp add: d'_def min_def) | |
| 60178 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 887 | fix x1 | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 888 |           assume x1: "x1 \<in> {a..<d'}"
 | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 889 |           {
 | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 890 | assume "x1 < y" | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 891 | hence "?le x1" | 
| 60420 | 892 |               using \<open>x1 \<in> {a..<d'}\<close> y_all_le by auto
 | 
| 60178 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 893 |           } moreover {
 | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 894 | assume "x1 \<ge> y" | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 895 |             hence x1': "x1 \<in> S" "x1 \<in> {y..<b}" using x1
 | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 896 | by (auto simp: d'_def dist_real_def intro!: d) | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 897 | have "norm (f x1 - f a) \<le> norm (f x1 - f y) + norm (f y - f a)" | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 898 | by (rule order_trans[OF _ norm_triangle_ineq]) simp | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 899 | also note S(3)[OF x1'] | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 900 | also note le_y | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 901 | finally have "?le x1" | 
| 60420 | 902 | using \<open>x1 \<ge> y\<close> by (auto simp: algebra_simps) | 
| 60178 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 903 | } ultimately show "?le x1" by arith | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 904 | qed | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 905 | hence "d' \<le> y" | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 906 | unfolding y_def | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 907 | by (rule cSup_upper) simp | 
| 60420 | 908 | thus False using \<open>d > 0\<close> \<open>y < b\<close> | 
| 62390 | 909 | by (simp add: d'_def min_def split: if_split_asm) | 
| 60178 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 910 | qed | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 911 |     } moreover {
 | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 912 | assume "a = y" | 
| 60420 | 913 | with \<open>a < b\<close> have "y < b" by simp | 
| 914 | with \<open>a = y\<close> f_cont phi_cont \<open>e2 > 0\<close> | |
| 60178 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 915 |       have 1: "\<forall>\<^sub>F x in at y within {y..b}. dist (f x) (f y) < e2"
 | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 916 |        and 2: "\<forall>\<^sub>F x in at y within {y..b}. dist (\<phi> x) (\<phi> y) < e2"
 | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 917 | by (auto simp: continuous_on_def tendsto_iff) | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 918 |       have 3: "eventually (\<lambda>x. y < x) (at y within {y..b})"
 | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 919 | by (auto simp: eventually_at_filter) | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 920 |       have 4: "eventually (\<lambda>x::real. x < b) (at y within {y..b})"
 | 
| 60420 | 921 | using _ \<open>y < b\<close> | 
| 60178 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 922 | by (rule order_tendstoD) (auto intro!: tendsto_eq_intros) | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 923 | from 1 2 3 4 | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 924 |       have eventually_le: "eventually (\<lambda>x. ?le x) (at y within {y .. b})"
 | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 925 | proof eventually_elim | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 926 | case (elim x1) | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 927 | have "norm (f x1 - f a) = norm (f x1 - f y)" | 
| 60420 | 928 | by (simp add: \<open>a = y\<close>) | 
| 60178 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 929 | also have "norm (f x1 - f y) \<le> e2" | 
| 60420 | 930 | using elim \<open>a = y\<close> by (auto simp : dist_norm intro!: less_imp_le) | 
| 60178 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 931 | also have "\<dots> \<le> e2 + (\<phi> x1 - \<phi> a + e2 + e * (x1 - a))" | 
| 60420 | 932 | using \<open>0 < e\<close> elim | 
| 60178 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 933 | by (intro add_increasing2[OF add_nonneg_nonneg order.refl]) | 
| 60420 | 934 | (auto simp: \<open>a = y\<close> dist_norm intro!: mult_nonneg_nonneg) | 
| 60178 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 935 | also have "\<dots> = \<phi> x1 - \<phi> a + e * (x1 - a) + e" | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 936 | by (simp add: e2_def) | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 937 | finally show "?le x1" . | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 938 | qed | 
| 60420 | 939 | from this[unfolded eventually_at_topological] \<open>?le y\<close> | 
| 60178 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 940 | obtain S | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 941 |       where S: "open S" "y \<in> S" "\<And>x. x\<in>S \<Longrightarrow> x \<in> {y..b} \<Longrightarrow> ?le x"
 | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 942 | by metis | 
| 60420 | 943 | from \<open>open S\<close> obtain d where d: "\<And>x. dist x y < d \<Longrightarrow> x \<in> S" "d > 0" | 
| 62101 | 944 | by (force simp: dist_commute open_dist ball_def dest!: bspec[OF _ \<open>y \<in> S\<close>]) | 
| 60178 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 945 | def d' \<equiv> "min b (y + (d/2))" | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 946 | have "d' \<in> A" | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 947 | unfolding A_def | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 948 | proof safe | 
| 60420 | 949 | show "a \<le> d'" using \<open>a = y\<close> \<open>0 < d\<close> \<open>y < b\<close> by (simp add: d'_def) | 
| 60178 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 950 | show "d' \<le> b" by (simp add: d'_def) | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 951 | fix x1 | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 952 |         assume "x1 \<in> {a..<d'}"
 | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 953 |         hence "x1 \<in> S" "x1 \<in> {y..b}"
 | 
| 60420 | 954 | by (auto simp: \<open>a = y\<close> d'_def dist_real_def intro!: d ) | 
| 60178 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 955 | thus "?le x1" | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 956 | by (rule S) | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 957 | qed | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 958 | hence "d' \<le> y" | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 959 | unfolding y_def | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 960 | by (rule cSup_upper) simp | 
| 60420 | 961 | hence "y = b" using \<open>d > 0\<close> \<open>y < b\<close> | 
| 60178 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 962 | by (simp add: d'_def) | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 963 | } ultimately have "y = b" | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 964 | by auto | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 965 | with le_y have "norm (f b - f a) \<le> \<phi> b - \<phi> a + e * (b - a + 1)" | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 966 | by (simp add: algebra_simps) | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 967 | } note * = this | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 968 |   {
 | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 969 | fix e::real assume "e > 0" | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 970 | hence "norm (f b - f a) \<le> \<phi> b - \<phi> a + e" | 
| 60420 | 971 | using *[of "e / (b - a + 1)"] \<open>a < b\<close> by simp | 
| 60178 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 972 | } thus ?thesis | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 973 | by (rule field_le_epsilon) | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 974 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 975 | |
| 44123 | 976 | lemma differentiable_bound: | 
| 60178 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 977 | fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" | 
| 53781 | 978 | assumes "convex s" | 
| 979 | and "\<forall>x\<in>s. (f has_derivative f' x) (at x within s)" | |
| 980 | and "\<forall>x\<in>s. onorm (f' x) \<le> B" | |
| 981 | and x: "x \<in> s" | |
| 982 | and y: "y \<in> s" | |
| 983 | shows "norm (f x - f y) \<le> B * norm (x - y)" | |
| 984 | proof - | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 985 | let ?p = "\<lambda>u. x + u *\<^sub>R (y - x)" | 
| 60178 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 986 | let ?\<phi> = "\<lambda>h. h * B * norm (x - y)" | 
| 53781 | 987 |   have *: "\<And>u. u\<in>{0..1} \<Longrightarrow> x + u *\<^sub>R (y - x) \<in> s"
 | 
| 44123 | 988 | using assms(1)[unfolded convex_alt,rule_format,OF x y] | 
| 989 | unfolding scaleR_left_diff_distrib scaleR_right_diff_distrib | |
| 990 | by (auto simp add: algebra_simps) | |
| 60178 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 991 |   have 0: "continuous_on (?p ` {0..1}) f"
 | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 992 | using * | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 993 | unfolding continuous_on_eq_continuous_within | 
| 53781 | 994 | apply - | 
| 995 | apply rule | |
| 996 | apply (rule differentiable_imp_continuous_within) | |
| 997 | unfolding differentiable_def | |
| 998 | apply (rule_tac x="f' xa" in exI) | |
| 999 | apply (rule has_derivative_within_subset) | |
| 1000 | apply (rule assms(2)[rule_format]) | |
| 1001 | apply auto | |
| 1002 | done | |
| 60178 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1003 |   from * have 1: "continuous_on {0 .. 1} (f \<circ> ?p)"
 | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1004 | by (intro continuous_intros 0)+ | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1005 |   {
 | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1006 |     fix u::real assume u: "u \<in>{0 <..< 1}"
 | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1007 | let ?u = "?p u" | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1008 | interpret linear "(f' ?u)" | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1009 | using u by (auto intro!: has_derivative_linear assms(2)[rule_format] *) | 
| 56188 | 1010 | have "(f \<circ> ?p has_derivative (f' ?u) \<circ> (\<lambda>u. 0 + u *\<^sub>R (y - x))) (at u within box 0 1)" | 
| 53781 | 1011 | apply (rule diff_chain_within) | 
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 1012 | apply (rule derivative_intros)+ | 
| 53781 | 1013 | apply (rule has_derivative_within_subset) | 
| 1014 | apply (rule assms(2)[rule_format]) | |
| 60178 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1015 | using u * | 
| 53781 | 1016 | apply auto | 
| 1017 | done | |
| 60178 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1018 | hence "((f \<circ> ?p) has_vector_derivative f' ?u (y - x)) (at u)" | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1019 | by (simp add: has_derivative_within_open[OF u open_greaterThanLessThan] | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1020 | scaleR has_vector_derivative_def o_def) | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1021 | } note 2 = this | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1022 |   {
 | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1023 |     have "continuous_on {0..1} ?\<phi>"
 | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1024 | by (rule continuous_intros)+ | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1025 | } note 3 = this | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1026 |   {
 | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1027 |     fix u::real assume u: "u \<in>{0 <..< 1}"
 | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1028 | have "(?\<phi> has_vector_derivative B * norm (x - y)) (at u)" | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1029 | by (auto simp: has_vector_derivative_def intro!: derivative_eq_intros) | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1030 | } note 4 = this | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1031 |   {
 | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1032 |     fix u::real assume u: "u \<in>{0 <..< 1}"
 | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1033 | let ?u = "?p u" | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1034 | interpret bounded_linear "(f' ?u)" | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1035 | using u by (auto intro!: has_derivative_bounded_linear assms(2)[rule_format] *) | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1036 | have "norm (f' ?u (y - x)) \<le> onorm (f' ?u) * norm (y - x)" | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1037 | by (rule onorm) fact | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1038 | also have "onorm (f' ?u) \<le> B" | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1039 | using u by (auto intro!: assms(3)[rule_format] *) | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1040 | finally have "norm ((f' ?u) (y - x)) \<le> B * norm (x - y)" | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1041 | by (simp add: mult_right_mono norm_minus_commute) | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1042 | } note 5 = this | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1043 | have "norm (f x - f y) = norm ((f \<circ> (\<lambda>u. x + u *\<^sub>R (y - x))) 1 - (f \<circ> (\<lambda>u. x + u *\<^sub>R (y - x))) 0)" | 
| 53781 | 1044 | by (auto simp add: norm_minus_commute) | 
| 60178 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1045 | also | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1046 | from differentiable_bound_general[OF zero_less_one 1, OF 3 2 4 5] | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1047 | have "norm ((f \<circ> ?p) 1 - (f \<circ> ?p) 0) \<le> B * norm (x - y)" | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1048 | by simp | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1049 | finally show ?thesis . | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1050 | qed | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1051 | |
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1052 | lemma | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1053 | differentiable_bound_segment: | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1054 | fixes f::"'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1055 |   assumes "\<And>t. t \<in> {0..1} \<Longrightarrow> x0 + t *\<^sub>R a \<in> G"
 | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1056 | assumes f': "\<And>x. x \<in> G \<Longrightarrow> (f has_derivative f' x) (at x within G)" | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1057 |   assumes B: "\<forall>x\<in>{0..1}. onorm (f' (x0 + x *\<^sub>R a)) \<le> B"
 | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1058 | shows "norm (f (x0 + a) - f x0) \<le> norm a * B" | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1059 | proof - | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1060 |   let ?G = "(\<lambda>x. x0 + x *\<^sub>R a) ` {0..1}"
 | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1061 |   have "?G = op + x0 ` (\<lambda>x. x *\<^sub>R a) ` {0..1}" by auto
 | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1062 | also have "convex \<dots>" | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1063 | by (intro convex_translation convex_scaled convex_real_interval) | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1064 | finally have "convex ?G" . | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1065 | moreover have "?G \<subseteq> G" "x0 \<in> ?G" "x0 + a \<in> ?G" using assms by (auto intro: image_eqI[where x=1]) | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1066 | ultimately show ?thesis | 
| 60420 | 1067 | using has_derivative_subset[OF f' \<open>?G \<subseteq> G\<close>] B | 
| 60178 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1068 |       differentiable_bound[of "(\<lambda>x. x0 + x *\<^sub>R a) ` {0..1}" f f' B "x0 + a" x0]
 | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1069 | by (auto simp: ac_simps) | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1070 | qed | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1071 | |
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1072 | lemma differentiable_bound_linearization: | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1073 | fixes f::"'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1074 |   assumes "\<And>t. t \<in> {0..1} \<Longrightarrow> a + t *\<^sub>R (b - a) \<in> S"
 | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1075 | assumes f'[derivative_intros]: "\<And>x. x \<in> S \<Longrightarrow> (f has_derivative f' x) (at x within S)" | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1076 | assumes B: "\<forall>x\<in>S. onorm (f' x - f' x0) \<le> B" | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1077 | assumes "x0 \<in> S" | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1078 | shows "norm (f b - f a - f' x0 (b - a)) \<le> norm (b - a) * B" | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1079 | proof - | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1080 | def g \<equiv> "\<lambda>x. f x - f' x0 x" | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1081 | have g: "\<And>x. x \<in> S \<Longrightarrow> (g has_derivative (\<lambda>i. f' x i - f' x0 i)) (at x within S)" | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1082 | unfolding g_def using assms | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1083 | by (auto intro!: derivative_eq_intros | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1084 | bounded_linear.has_derivative[OF has_derivative_bounded_linear, OF f']) | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1085 |   from B have B: "\<forall>x\<in>{0..1}. onorm (\<lambda>i. f' (a + x *\<^sub>R (b - a)) i - f' x0 i) \<le> B"
 | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1086 | using assms by (auto simp: fun_diff_def) | 
| 60420 | 1087 | from differentiable_bound_segment[OF assms(1) g B] \<open>x0 \<in> S\<close> | 
| 60178 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1088 | show ?thesis | 
| 
f620c70f9e9b
generalized differentiable_bound; some further variations of differentiable_bound
 immler parents: 
60177diff
changeset | 1089 | by (simp add: g_def field_simps linear_sub[OF has_derivative_linear[OF f']]) | 
| 44123 | 1090 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1091 | |
| 60420 | 1092 | text \<open>In particular.\<close> | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1093 | |
| 44123 | 1094 | lemma has_derivative_zero_constant: | 
| 60179 | 1095 | fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" | 
| 53781 | 1096 | assumes "convex s" | 
| 56369 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 hoelzl parents: 
56332diff
changeset | 1097 | and "\<And>x. x \<in> s \<Longrightarrow> (f has_derivative (\<lambda>h. 0)) (at x within s)" | 
| 44123 | 1098 | shows "\<exists>c. \<forall>x\<in>s. f x = c" | 
| 56332 | 1099 | proof - | 
| 1100 |   { fix x y assume "x \<in> s" "y \<in> s"
 | |
| 1101 | then have "norm (f x - f y) \<le> 0 * norm (x - y)" | |
| 1102 | using assms by (intro differentiable_bound[of s]) (auto simp: onorm_zero) | |
| 1103 | then have "f x = f y" | |
| 1104 | by simp } | |
| 53781 | 1105 | then show ?thesis | 
| 56332 | 1106 | by metis | 
| 53781 | 1107 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1108 | |
| 61524 
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
 eberlm parents: 
61520diff
changeset | 1109 | lemma has_field_derivative_zero_constant: | 
| 
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
 eberlm parents: 
61520diff
changeset | 1110 | assumes "convex s" "\<And>x. x \<in> s \<Longrightarrow> (f has_field_derivative 0) (at x within s)" | 
| 
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
 eberlm parents: 
61520diff
changeset | 1111 | shows "\<exists>c. \<forall>x\<in>s. f (x) = (c :: 'a :: real_normed_field)" | 
| 
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
 eberlm parents: 
61520diff
changeset | 1112 | proof (rule has_derivative_zero_constant) | 
| 
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
 eberlm parents: 
61520diff
changeset | 1113 | have A: "op * 0 = (\<lambda>_. 0 :: 'a)" by (intro ext) simp | 
| 
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
 eberlm parents: 
61520diff
changeset | 1114 | fix x assume "x \<in> s" thus "(f has_derivative (\<lambda>h. 0)) (at x within s)" | 
| 
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
 eberlm parents: 
61520diff
changeset | 1115 | using assms(2)[of x] by (simp add: has_field_derivative_def A) | 
| 
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
 eberlm parents: 
61520diff
changeset | 1116 | qed fact | 
| 
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
 eberlm parents: 
61520diff
changeset | 1117 | |
| 53781 | 1118 | lemma has_derivative_zero_unique: | 
| 60179 | 1119 | fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" | 
| 53781 | 1120 | assumes "convex s" | 
| 56369 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 hoelzl parents: 
56332diff
changeset | 1121 | and "\<And>x. x \<in> s \<Longrightarrow> (f has_derivative (\<lambda>h. 0)) (at x within s)" | 
| 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 hoelzl parents: 
56332diff
changeset | 1122 | and "x \<in> s" "y \<in> s" | 
| 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 hoelzl parents: 
56332diff
changeset | 1123 | shows "f x = f y" | 
| 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 hoelzl parents: 
56332diff
changeset | 1124 | using has_derivative_zero_constant[OF assms(1,2)] assms(3-) by force | 
| 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 hoelzl parents: 
56332diff
changeset | 1125 | |
| 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 hoelzl parents: 
56332diff
changeset | 1126 | lemma has_derivative_zero_unique_connected: | 
| 60179 | 1127 | fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" | 
| 56369 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 hoelzl parents: 
56332diff
changeset | 1128 | assumes "open s" "connected s" | 
| 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 hoelzl parents: 
56332diff
changeset | 1129 | assumes f: "\<And>x. x \<in> s \<Longrightarrow> (f has_derivative (\<lambda>x. 0)) (at x)" | 
| 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 hoelzl parents: 
56332diff
changeset | 1130 | assumes "x \<in> s" "y \<in> s" | 
| 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 hoelzl parents: 
56332diff
changeset | 1131 | shows "f x = f y" | 
| 60420 | 1132 | proof (rule connected_local_const[where f=f, OF \<open>connected s\<close> \<open>x\<in>s\<close> \<open>y\<in>s\<close>]) | 
| 56369 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 hoelzl parents: 
56332diff
changeset | 1133 | show "\<forall>a\<in>s. eventually (\<lambda>b. f a = f b) (at a within s)" | 
| 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 hoelzl parents: 
56332diff
changeset | 1134 | proof | 
| 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 hoelzl parents: 
56332diff
changeset | 1135 | fix a assume "a \<in> s" | 
| 60420 | 1136 | with \<open>open s\<close> obtain e where "0 < e" "ball a e \<subseteq> s" | 
| 56369 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 hoelzl parents: 
56332diff
changeset | 1137 | by (rule openE) | 
| 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 hoelzl parents: 
56332diff
changeset | 1138 | then have "\<exists>c. \<forall>x\<in>ball a e. f x = c" | 
| 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 hoelzl parents: 
56332diff
changeset | 1139 | by (intro has_derivative_zero_constant) | 
| 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 hoelzl parents: 
56332diff
changeset | 1140 | (auto simp: at_within_open[OF _ open_ball] f convex_ball) | 
| 60420 | 1141 | with \<open>0<e\<close> have "\<forall>x\<in>ball a e. f a = f x" | 
| 56369 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 hoelzl parents: 
56332diff
changeset | 1142 | by auto | 
| 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 hoelzl parents: 
56332diff
changeset | 1143 | then show "eventually (\<lambda>b. f a = f b) (at a within s)" | 
| 60420 | 1144 | using \<open>0<e\<close> unfolding eventually_at_topological | 
| 56369 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 hoelzl parents: 
56332diff
changeset | 1145 | by (intro exI[of _ "ball a e"]) auto | 
| 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 hoelzl parents: 
56332diff
changeset | 1146 | qed | 
| 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 hoelzl parents: 
56332diff
changeset | 1147 | qed | 
| 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 hoelzl parents: 
56332diff
changeset | 1148 | |
| 60420 | 1149 | subsection \<open>Differentiability of inverse function (most basic form)\<close> | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1150 | |
| 44123 | 1151 | lemma has_derivative_inverse_basic: | 
| 56226 | 1152 | fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" | 
| 44123 | 1153 | assumes "(f has_derivative f') (at (g y))" | 
| 53781 | 1154 | and "bounded_linear g'" | 
| 1155 | and "g' \<circ> f' = id" | |
| 1156 | and "continuous (at y) g" | |
| 1157 | and "open t" | |
| 1158 | and "y \<in> t" | |
| 1159 | and "\<forall>z\<in>t. f (g z) = z" | |
| 44123 | 1160 | shows "(g has_derivative g') (at y)" | 
| 53781 | 1161 | proof - | 
| 44123 | 1162 | interpret f': bounded_linear f' | 
| 1163 | using assms unfolding has_derivative_def by auto | |
| 53781 | 1164 | interpret g': bounded_linear g' | 
| 1165 | using assms by auto | |
| 55665 | 1166 | obtain C where C: "0 < C" "\<And>x. norm (g' x) \<le> norm x * C" | 
| 1167 | using bounded_linear.pos_bounded[OF assms(2)] by blast | |
| 53781 | 1168 | have lem1: "\<forall>e>0. \<exists>d>0. \<forall>z. | 
| 1169 | norm (z - y) < d \<longrightarrow> norm (g z - g y - g'(z - y)) \<le> e * norm (g z - g y)" | |
| 1170 | proof (rule, rule) | |
| 61165 | 1171 | fix e :: real | 
| 1172 | assume "e > 0" | |
| 1173 | with C(1) have *: "e / C > 0" by auto | |
| 55665 | 1174 | obtain d0 where d0: | 
| 1175 | "0 < d0" | |
| 1176 | "\<forall>ya. norm (ya - g y) < d0 \<longrightarrow> norm (f ya - f (g y) - f' (ya - g y)) \<le> e / C * norm (ya - g y)" | |
| 1177 | using assms(1) | |
| 1178 | unfolding has_derivative_at_alt | |
| 1179 | using * by blast | |
| 1180 | obtain d1 where d1: | |
| 1181 | "0 < d1" | |
| 1182 | "\<forall>x. 0 < dist x y \<and> dist x y < d1 \<longrightarrow> dist (g x) (g y) < d0" | |
| 1183 | using assms(4) | |
| 1184 | unfolding continuous_at Lim_at | |
| 1185 | using d0(1) by blast | |
| 1186 | obtain d2 where d2: | |
| 1187 | "0 < d2" | |
| 1188 | "\<forall>ya. dist ya y < d2 \<longrightarrow> ya \<in> t" | |
| 1189 | using assms(5) | |
| 1190 | unfolding open_dist | |
| 1191 | using assms(6) by blast | |
| 1192 | obtain d where d: "0 < d" "d < d1" "d < d2" | |
| 1193 | using real_lbound_gt_zero[OF d1(1) d2(1)] by blast | |
| 61165 | 1194 | then show "\<exists>d>0. \<forall>z. norm (z - y) < d \<longrightarrow> norm (g z - g y - g' (z - y)) \<le> e * norm (g z - g y)" | 
| 53781 | 1195 | apply (rule_tac x=d in exI) | 
| 1196 | apply rule | |
| 1197 | defer | |
| 1198 | apply rule | |
| 1199 | apply rule | |
| 1200 | proof - | |
| 1201 | fix z | |
| 1202 | assume as: "norm (z - y) < d" | |
| 1203 | then have "z \<in> t" | |
| 44123 | 1204 | using d2 d unfolding dist_norm by auto | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1205 | have "norm (g z - g y - g' (z - y)) \<le> norm (g' (f (g z) - y - f' (g z - g y)))" | 
| 44123 | 1206 | unfolding g'.diff f'.diff | 
| 53781 | 1207 | unfolding assms(3)[unfolded o_def id_def, THEN fun_cong] | 
| 60420 | 1208 | unfolding assms(7)[rule_format,OF \<open>z\<in>t\<close>] | 
| 53781 | 1209 | apply (subst norm_minus_cancel[symmetric]) | 
| 1210 | apply auto | |
| 1211 | done | |
| 1212 | also have "\<dots> \<le> norm (f (g z) - y - f' (g z - g y)) * C" | |
| 55665 | 1213 | by (rule C(2)) | 
| 44123 | 1214 | also have "\<dots> \<le> (e / C) * norm (g z - g y) * C" | 
| 53781 | 1215 | apply (rule mult_right_mono) | 
| 60420 | 1216 | apply (rule d0(2)[rule_format,unfolded assms(7)[rule_format,OF \<open>y\<in>t\<close>]]) | 
| 53781 | 1217 | apply (cases "z = y") | 
| 1218 | defer | |
| 55665 | 1219 | apply (rule d1(2)[unfolded dist_norm,rule_format]) | 
| 53781 | 1220 | using as d C d0 | 
| 1221 | apply auto | |
| 1222 | done | |
| 44123 | 1223 | also have "\<dots> \<le> e * norm (g z - g y)" | 
| 1224 | using C by (auto simp add: field_simps) | |
| 1225 | finally show "norm (g z - g y - g' (z - y)) \<le> e * norm (g z - g y)" | |
| 1226 | by simp | |
| 1227 | qed auto | |
| 1228 | qed | |
| 53781 | 1229 | have *: "(0::real) < 1 / 2" | 
| 1230 | by auto | |
| 55665 | 1231 | obtain d where d: | 
| 1232 | "0 < d" | |
| 1233 | "\<forall>z. norm (z - y) < d \<longrightarrow> norm (g z - g y - g' (z - y)) \<le> 1 / 2 * norm (g z - g y)" | |
| 1234 | using lem1 * by blast | |
| 1235 | def B \<equiv> "C * 2" | |
| 53781 | 1236 | have "B > 0" | 
| 1237 | unfolding B_def using C by auto | |
| 61165 | 1238 | have lem2: "norm (g z - g y) \<le> B * norm (z - y)" if z: "norm(z - y) < d" for z | 
| 1239 | proof - | |
| 44123 | 1240 | have "norm (g z - g y) \<le> norm(g' (z - y)) + norm ((g z - g y) - g'(z - y))" | 
| 53781 | 1241 | by (rule norm_triangle_sub) | 
| 1242 | also have "\<dots> \<le> norm (g' (z - y)) + 1 / 2 * norm (g z - g y)" | |
| 1243 | apply (rule add_left_mono) | |
| 61165 | 1244 | using d and z | 
| 53781 | 1245 | apply auto | 
| 1246 | done | |
| 44123 | 1247 | also have "\<dots> \<le> norm (z - y) * C + 1 / 2 * norm (g z - g y)" | 
| 53781 | 1248 | apply (rule add_right_mono) | 
| 1249 | using C | |
| 1250 | apply auto | |
| 1251 | done | |
| 61165 | 1252 | finally show "norm (g z - g y) \<le> B * norm (z - y)" | 
| 53781 | 1253 | unfolding B_def | 
| 1254 | by (auto simp add: field_simps) | |
| 44123 | 1255 | qed | 
| 53781 | 1256 | show ?thesis | 
| 1257 | unfolding has_derivative_at_alt | |
| 1258 | apply rule | |
| 1259 | apply (rule assms) | |
| 1260 | apply rule | |
| 1261 | apply rule | |
| 1262 | proof - | |
| 61165 | 1263 | fix e :: real | 
| 1264 | assume "e > 0" | |
| 1265 | then have *: "e / B > 0" by (metis \<open>B > 0\<close> divide_pos_pos) | |
| 55665 | 1266 | obtain d' where d': | 
| 1267 | "0 < d'" | |
| 1268 | "\<forall>z. norm (z - y) < d' \<longrightarrow> norm (g z - g y - g' (z - y)) \<le> e / B * norm (g z - g y)" | |
| 1269 | using lem1 * by blast | |
| 1270 | obtain k where k: "0 < k" "k < d" "k < d'" | |
| 1271 | using real_lbound_gt_zero[OF d(1) d'(1)] by blast | |
| 61165 | 1272 | show "\<exists>d>0. \<forall>ya. norm (ya - y) < d \<longrightarrow> norm (g ya - g y - g' (ya - y)) \<le> e * norm (ya - y)" | 
| 53781 | 1273 | apply (rule_tac x=k in exI) | 
| 55970 
6d123f0ae358
Some new proofs. Tidying up, esp to remove "apply rule".
 paulson <lp15@cam.ac.uk> parents: 
55665diff
changeset | 1274 | apply auto | 
| 53781 | 1275 | proof - | 
| 1276 | fix z | |
| 1277 | assume as: "norm (z - y) < k" | |
| 1278 | then have "norm (g z - g y - g' (z - y)) \<le> e / B * norm(g z - g y)" | |
| 44123 | 1279 | using d' k by auto | 
| 53781 | 1280 | also have "\<dots> \<le> e * norm (z - y)" | 
| 60420 | 1281 | unfolding times_divide_eq_left pos_divide_le_eq[OF \<open>B>0\<close>] | 
| 61165 | 1282 | using lem2[of z] | 
| 60420 | 1283 | using k as using \<open>e > 0\<close> | 
| 44123 | 1284 | by (auto simp add: field_simps) | 
| 1285 | finally show "norm (g z - g y - g' (z - y)) \<le> e * norm (z - y)" | |
| 53781 | 1286 | by simp | 
| 1287 | qed(insert k, auto) | |
| 44123 | 1288 | qed | 
| 1289 | qed | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1290 | |
| 60420 | 1291 | text \<open>Simply rewrite that based on the domain point x.\<close> | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1292 | |
| 44123 | 1293 | lemma has_derivative_inverse_basic_x: | 
| 56226 | 1294 | fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" | 
| 53781 | 1295 | assumes "(f has_derivative f') (at x)" | 
| 1296 | and "bounded_linear g'" | |
| 1297 | and "g' \<circ> f' = id" | |
| 1298 | and "continuous (at (f x)) g" | |
| 1299 | and "g (f x) = x" | |
| 1300 | and "open t" | |
| 1301 | and "f x \<in> t" | |
| 1302 | and "\<forall>y\<in>t. f (g y) = y" | |
| 1303 | shows "(g has_derivative g') (at (f x))" | |
| 1304 | apply (rule has_derivative_inverse_basic) | |
| 1305 | using assms | |
| 1306 | apply auto | |
| 1307 | done | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1308 | |
| 60420 | 1309 | text \<open>This is the version in Dieudonne', assuming continuity of f and g.\<close> | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1310 | |
| 44123 | 1311 | lemma has_derivative_inverse_dieudonne: | 
| 56226 | 1312 | fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" | 
| 53781 | 1313 | assumes "open s" | 
| 1314 | and "open (f ` s)" | |
| 1315 | and "continuous_on s f" | |
| 1316 | and "continuous_on (f ` s) g" | |
| 1317 | and "\<forall>x\<in>s. g (f x) = x" | |
| 1318 | and "x \<in> s" | |
| 1319 | and "(f has_derivative f') (at x)" | |
| 1320 | and "bounded_linear g'" | |
| 1321 | and "g' \<circ> f' = id" | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1322 | shows "(g has_derivative g') (at (f x))" | 
| 53781 | 1323 | apply (rule has_derivative_inverse_basic_x[OF assms(7-9) _ _ assms(2)]) | 
| 1324 | using assms(3-6) | |
| 1325 | unfolding continuous_on_eq_continuous_at[OF assms(1)] continuous_on_eq_continuous_at[OF assms(2)] | |
| 1326 | apply auto | |
| 1327 | done | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1328 | |
| 60420 | 1329 | text \<open>Here's the simplest way of not assuming much about g.\<close> | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1330 | |
| 44123 | 1331 | lemma has_derivative_inverse: | 
| 56226 | 1332 | fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" | 
| 53781 | 1333 | assumes "compact s" | 
| 1334 | and "x \<in> s" | |
| 1335 | and "f x \<in> interior (f ` s)" | |
| 1336 | and "continuous_on s f" | |
| 1337 | and "\<forall>y\<in>s. g (f y) = y" | |
| 1338 | and "(f has_derivative f') (at x)" | |
| 1339 | and "bounded_linear g'" | |
| 1340 | and "g' \<circ> f' = id" | |
| 44123 | 1341 | shows "(g has_derivative g') (at (f x))" | 
| 53781 | 1342 | proof - | 
| 1343 |   {
 | |
| 1344 | fix y | |
| 1345 | assume "y \<in> interior (f ` s)" | |
| 1346 | then obtain x where "x \<in> s" and *: "y = f x" | |
| 1347 | unfolding image_iff | |
| 1348 | using interior_subset | |
| 1349 | by auto | |
| 1350 | have "f (g y) = y" | |
| 60420 | 1351 | unfolding * and assms(5)[rule_format,OF \<open>x\<in>s\<close>] .. | 
| 44123 | 1352 | } note * = this | 
| 1353 | show ?thesis | |
| 53781 | 1354 | apply (rule has_derivative_inverse_basic_x[OF assms(6-8)]) | 
| 1355 | apply (rule continuous_on_interior[OF _ assms(3)]) | |
| 1356 | apply (rule continuous_on_inv[OF assms(4,1)]) | |
| 1357 | apply (rule assms(2,5) assms(5)[rule_format] open_interior assms(3))+ | |
| 55970 
6d123f0ae358
Some new proofs. Tidying up, esp to remove "apply rule".
 paulson <lp15@cam.ac.uk> parents: 
55665diff
changeset | 1358 | apply (metis *) | 
| 53781 | 1359 | done | 
| 44123 | 1360 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1361 | |
| 53781 | 1362 | |
| 60420 | 1363 | subsection \<open>Proving surjectivity via Brouwer fixpoint theorem\<close> | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1364 | |
| 44123 | 1365 | lemma brouwer_surjective: | 
| 56117 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55970diff
changeset | 1366 | fixes f :: "'n::euclidean_space \<Rightarrow> 'n" | 
| 53781 | 1367 | assumes "compact t" | 
| 1368 | and "convex t" | |
| 1369 |     and "t \<noteq> {}"
 | |
| 1370 | and "continuous_on t f" | |
| 1371 | and "\<forall>x\<in>s. \<forall>y\<in>t. x + (y - f y) \<in> t" | |
| 1372 | and "x \<in> s" | |
| 44123 | 1373 | shows "\<exists>y\<in>t. f y = x" | 
| 53781 | 1374 | proof - | 
| 1375 | have *: "\<And>x y. f y = x \<longleftrightarrow> x + (y - f y) = y" | |
| 1376 | by (auto simp add: algebra_simps) | |
| 44123 | 1377 | show ?thesis | 
| 1378 | unfolding * | |
| 53781 | 1379 | apply (rule brouwer[OF assms(1-3), of "\<lambda>y. x + (y - f y)"]) | 
| 56371 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 hoelzl parents: 
56370diff
changeset | 1380 | apply (rule continuous_intros assms)+ | 
| 53781 | 1381 | using assms(4-6) | 
| 1382 | apply auto | |
| 1383 | done | |
| 44123 | 1384 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1385 | |
| 44123 | 1386 | lemma brouwer_surjective_cball: | 
| 56117 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55970diff
changeset | 1387 | fixes f :: "'n::euclidean_space \<Rightarrow> 'n" | 
| 53781 | 1388 | assumes "e > 0" | 
| 1389 | and "continuous_on (cball a e) f" | |
| 1390 | and "\<forall>x\<in>s. \<forall>y\<in>cball a e. x + (y - f y) \<in> cball a e" | |
| 1391 | and "x \<in> s" | |
| 44123 | 1392 | shows "\<exists>y\<in>cball a e. f y = x" | 
| 53781 | 1393 | apply (rule brouwer_surjective) | 
| 1394 | apply (rule compact_cball convex_cball)+ | |
| 1395 | unfolding cball_eq_empty | |
| 1396 | using assms | |
| 1397 | apply auto | |
| 1398 | done | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1399 | |
| 60420 | 1400 | text \<open>See Sussmann: "Multidifferential calculus", Theorem 2.1.1\<close> | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1401 | |
| 44123 | 1402 | lemma sussmann_open_mapping: | 
| 56227 | 1403 | fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::euclidean_space" | 
| 53781 | 1404 | assumes "open s" | 
| 1405 | and "continuous_on s f" | |
| 1406 | and "x \<in> s" | |
| 1407 | and "(f has_derivative f') (at x)" | |
| 1408 | and "bounded_linear g'" "f' \<circ> g' = id" | |
| 1409 | and "t \<subseteq> s" | |
| 1410 | and "x \<in> interior t" | |
| 44123 | 1411 | shows "f x \<in> interior (f ` t)" | 
| 53781 | 1412 | proof - | 
| 1413 | interpret f': bounded_linear f' | |
| 1414 | using assms | |
| 1415 | unfolding has_derivative_def | |
| 1416 | by auto | |
| 1417 | interpret g': bounded_linear g' | |
| 1418 | using assms | |
| 1419 | by auto | |
| 55665 | 1420 | obtain B where B: "0 < B" "\<forall>x. norm (g' x) \<le> norm x * B" | 
| 1421 | using bounded_linear.pos_bounded[OF assms(5)] by blast | |
| 56541 | 1422 | hence *: "1 / (2 * B) > 0" by auto | 
| 55665 | 1423 | obtain e0 where e0: | 
| 1424 | "0 < e0" | |
| 1425 | "\<forall>y. norm (y - x) < e0 \<longrightarrow> norm (f y - f x - f' (y - x)) \<le> 1 / (2 * B) * norm (y - x)" | |
| 1426 | using assms(4) | |
| 1427 | unfolding has_derivative_at_alt | |
| 1428 | using * by blast | |
| 1429 | obtain e1 where e1: "0 < e1" "cball x e1 \<subseteq> t" | |
| 1430 | using assms(8) | |
| 1431 | unfolding mem_interior_cball | |
| 1432 | by blast | |
| 56541 | 1433 | have *: "0 < e0 / B" "0 < e1 / B" using e0 e1 B by auto | 
| 55665 | 1434 | obtain e where e: "0 < e" "e < e0 / B" "e < e1 / B" | 
| 1435 | using real_lbound_gt_zero[OF *] by blast | |
| 53781 | 1436 | have "\<forall>z\<in>cball (f x) (e / 2). \<exists>y\<in>cball (f x) e. f (x + g' (y - f x)) = z" | 
| 1437 | apply rule | |
| 1438 | apply (rule brouwer_surjective_cball[where s="cball (f x) (e/2)"]) | |
| 1439 | prefer 3 | |
| 1440 | apply rule | |
| 1441 | apply rule | |
| 44123 | 1442 | proof- | 
| 1443 | show "continuous_on (cball (f x) e) (\<lambda>y. f (x + g' (y - f x)))" | |
| 1444 | unfolding g'.diff | |
| 53781 | 1445 | apply (rule continuous_on_compose[of _ _ f, unfolded o_def]) | 
| 56371 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 hoelzl parents: 
56370diff
changeset | 1446 | apply (rule continuous_intros linear_continuous_on[OF assms(5)])+ | 
| 53781 | 1447 | apply (rule continuous_on_subset[OF assms(2)]) | 
| 1448 | apply rule | |
| 1449 | apply (unfold image_iff) | |
| 1450 | apply (erule bexE) | |
| 44123 | 1451 | proof- | 
| 53781 | 1452 | fix y z | 
| 1453 | assume as: "y \<in>cball (f x) e" "z = x + (g' y - g' (f x))" | |
| 44123 | 1454 | have "dist x z = norm (g' (f x) - g' y)" | 
| 1455 | unfolding as(2) and dist_norm by auto | |
| 1456 | also have "\<dots> \<le> norm (f x - y) * B" | |
| 53781 | 1457 | unfolding g'.diff[symmetric] | 
| 1458 | using B | |
| 1459 | by auto | |
| 44123 | 1460 | also have "\<dots> \<le> e * B" | 
| 53781 | 1461 | using as(1)[unfolded mem_cball dist_norm] | 
| 1462 | using B | |
| 1463 | by auto | |
| 1464 | also have "\<dots> \<le> e1" | |
| 1465 | using e | |
| 1466 | unfolding less_divide_eq | |
| 1467 | using B | |
| 1468 | by auto | |
| 1469 | finally have "z \<in> cball x e1" | |
| 1470 | unfolding mem_cball | |
| 1471 | by force | |
| 1472 | then show "z \<in> s" | |
| 1473 | using e1 assms(7) by auto | |
| 44123 | 1474 | qed | 
| 1475 | next | |
| 53781 | 1476 | fix y z | 
| 1477 | assume as: "y \<in> cball (f x) (e / 2)" "z \<in> cball (f x) e" | |
| 1478 | have "norm (g' (z - f x)) \<le> norm (z - f x) * B" | |
| 1479 | using B by auto | |
| 1480 | also have "\<dots> \<le> e * B" | |
| 1481 | apply (rule mult_right_mono) | |
| 44123 | 1482 | using as(2)[unfolded mem_cball dist_norm] and B | 
| 53781 | 1483 | unfolding norm_minus_commute | 
| 1484 | apply auto | |
| 1485 | done | |
| 1486 | also have "\<dots> < e0" | |
| 1487 | using e and B | |
| 1488 | unfolding less_divide_eq | |
| 1489 | by auto | |
| 1490 | finally have *: "norm (x + g' (z - f x) - x) < e0" | |
| 1491 | by auto | |
| 1492 | have **: "f x + f' (x + g' (z - f x) - x) = z" | |
| 1493 | using assms(6)[unfolded o_def id_def,THEN cong] | |
| 1494 | by auto | |
| 1495 | have "norm (f x - (y + (z - f (x + g' (z - f x))))) \<le> | |
| 1496 | norm (f (x + g' (z - f x)) - z) + norm (f x - y)" | |
| 44123 | 1497 | using norm_triangle_ineq[of "f (x + g'(z - f x)) - z" "f x - y"] | 
| 1498 | by (auto simp add: algebra_simps) | |
| 1499 | also have "\<dots> \<le> 1 / (B * 2) * norm (g' (z - f x)) + norm (f x - y)" | |
| 55665 | 1500 | using e0(2)[rule_format, OF *] | 
| 53781 | 1501 | unfolding algebra_simps ** | 
| 1502 | by auto | |
| 44123 | 1503 | also have "\<dots> \<le> 1 / (B * 2) * norm (g' (z - f x)) + e/2" | 
| 53781 | 1504 | using as(1)[unfolded mem_cball dist_norm] | 
| 1505 | by auto | |
| 44123 | 1506 | also have "\<dots> \<le> 1 / (B * 2) * B * norm (z - f x) + e/2" | 
| 53781 | 1507 | using * and B | 
| 1508 | by (auto simp add: field_simps) | |
| 1509 | also have "\<dots> \<le> 1 / 2 * norm (z - f x) + e/2" | |
| 1510 | by auto | |
| 1511 | also have "\<dots> \<le> e/2 + e/2" | |
| 1512 | apply (rule add_right_mono) | |
| 44123 | 1513 | using as(2)[unfolded mem_cball dist_norm] | 
| 53781 | 1514 | unfolding norm_minus_commute | 
| 1515 | apply auto | |
| 1516 | done | |
| 44123 | 1517 | finally show "y + (z - f (x + g' (z - f x))) \<in> cball (f x) e" | 
| 53781 | 1518 | unfolding mem_cball dist_norm | 
| 1519 | by auto | |
| 1520 | qed (insert e, auto) note lem = this | |
| 1521 | show ?thesis | |
| 1522 | unfolding mem_interior | |
| 1523 | apply (rule_tac x="e/2" in exI) | |
| 1524 | apply rule | |
| 1525 | apply (rule divide_pos_pos) | |
| 1526 | prefer 3 | |
| 44123 | 1527 | proof | 
| 53781 | 1528 | fix y | 
| 1529 | assume "y \<in> ball (f x) (e / 2)" | |
| 1530 | then have *: "y \<in> cball (f x) (e / 2)" | |
| 1531 | by auto | |
| 55665 | 1532 | obtain z where z: "z \<in> cball (f x) e" "f (x + g' (z - f x)) = y" | 
| 1533 | using lem * by blast | |
| 53781 | 1534 | then have "norm (g' (z - f x)) \<le> norm (z - f x) * B" | 
| 1535 | using B | |
| 1536 | by (auto simp add: field_simps) | |
| 44123 | 1537 | also have "\<dots> \<le> e * B" | 
| 53781 | 1538 | apply (rule mult_right_mono) | 
| 1539 | using z(1) | |
| 1540 | unfolding mem_cball dist_norm norm_minus_commute | |
| 1541 | using B | |
| 1542 | apply auto | |
| 1543 | done | |
| 1544 | also have "\<dots> \<le> e1" | |
| 1545 | using e B unfolding less_divide_eq by auto | |
| 1546 | finally have "x + g'(z - f x) \<in> t" | |
| 1547 | apply - | |
| 55665 | 1548 | apply (rule e1(2)[unfolded subset_eq,rule_format]) | 
| 53781 | 1549 | unfolding mem_cball dist_norm | 
| 1550 | apply auto | |
| 1551 | done | |
| 1552 | then show "y \<in> f ` t" | |
| 1553 | using z by auto | |
| 1554 | qed (insert e, auto) | |
| 44123 | 1555 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1556 | |
| 60420 | 1557 | text \<open>Hence the following eccentric variant of the inverse function theorem. | 
| 53799 | 1558 | This has no continuity assumptions, but we do need the inverse function. | 
| 61808 | 1559 | We could put \<open>f' \<circ> g = I\<close> but this happens to fit with the minimal linear | 
| 60420 | 1560 | algebra theory I've set up so far.\<close> | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1561 | |
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36844diff
changeset | 1562 | (* move before left_inverse_linear in Euclidean_Space*) | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36844diff
changeset | 1563 | |
| 53781 | 1564 | lemma right_inverse_linear: | 
| 1565 | fixes f :: "'a::euclidean_space \<Rightarrow> 'a" | |
| 1566 | assumes lf: "linear f" | |
| 1567 | and gf: "f \<circ> g = id" | |
| 1568 | shows "linear g" | |
| 1569 | proof - | |
| 1570 | from gf have fi: "surj f" | |
| 1571 | by (auto simp add: surj_def o_def id_def) metis | |
| 1572 | from linear_surjective_isomorphism[OF lf fi] | |
| 1573 | obtain h:: "'a \<Rightarrow> 'a" where h: "linear h" "\<forall>x. h (f x) = x" "\<forall>x. f (h x) = x" | |
| 1574 | by blast | |
| 1575 | have "h = g" | |
| 1576 | apply (rule ext) | |
| 1577 | using gf h(2,3) | |
| 1578 | apply (simp add: o_def id_def fun_eq_iff) | |
| 1579 | apply metis | |
| 1580 | done | |
| 1581 | with h(1) show ?thesis by blast | |
| 1582 | qed | |
| 1583 | ||
| 44123 | 1584 | lemma has_derivative_inverse_strong: | 
| 56117 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55970diff
changeset | 1585 | fixes f :: "'n::euclidean_space \<Rightarrow> 'n" | 
| 53781 | 1586 | assumes "open s" | 
| 1587 | and "x \<in> s" | |
| 1588 | and "continuous_on s f" | |
| 1589 | and "\<forall>x\<in>s. g (f x) = x" | |
| 1590 | and "(f has_derivative f') (at x)" | |
| 1591 | and "f' \<circ> g' = id" | |
| 44123 | 1592 | shows "(g has_derivative g') (at (f x))" | 
| 53781 | 1593 | proof - | 
| 1594 | have linf: "bounded_linear f'" | |
| 44123 | 1595 | using assms(5) unfolding has_derivative_def by auto | 
| 53781 | 1596 | then have ling: "bounded_linear g'" | 
| 1597 | unfolding linear_conv_bounded_linear[symmetric] | |
| 1598 | apply - | |
| 1599 | apply (rule right_inverse_linear) | |
| 1600 | using assms(6) | |
| 1601 | apply auto | |
| 1602 | done | |
| 1603 | moreover have "g' \<circ> f' = id" | |
| 1604 | using assms(6) linf ling | |
| 1605 | unfolding linear_conv_bounded_linear[symmetric] | |
| 1606 | using linear_inverse_left | |
| 1607 | by auto | |
| 1608 | moreover have *:"\<forall>t\<subseteq>s. x \<in> interior t \<longrightarrow> f x \<in> interior (f ` t)" | |
| 55970 
6d123f0ae358
Some new proofs. Tidying up, esp to remove "apply rule".
 paulson <lp15@cam.ac.uk> parents: 
55665diff
changeset | 1609 | apply clarify | 
| 53781 | 1610 | apply (rule sussmann_open_mapping) | 
| 1611 | apply (rule assms ling)+ | |
| 1612 | apply auto | |
| 1613 | done | |
| 1614 | have "continuous (at (f x)) g" | |
| 1615 | unfolding continuous_at Lim_at | |
| 1616 | proof (rule, rule) | |
| 1617 | fix e :: real | |
| 1618 | assume "e > 0" | |
| 1619 | then have "f x \<in> interior (f ` (ball x e \<inter> s))" | |
| 60420 | 1620 | using *[rule_format,of "ball x e \<inter> s"] \<open>x \<in> s\<close> | 
| 53781 | 1621 | by (auto simp add: interior_open[OF open_ball] interior_open[OF assms(1)]) | 
| 55665 | 1622 | then obtain d where d: "0 < d" "ball (f x) d \<subseteq> f ` (ball x e \<inter> s)" | 
| 1623 | unfolding mem_interior by blast | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1624 | show "\<exists>d>0. \<forall>y. 0 < dist y (f x) \<and> dist y (f x) < d \<longrightarrow> dist (g y) (g (f x)) < e" | 
| 53781 | 1625 | apply (rule_tac x=d in exI) | 
| 1626 | apply rule | |
| 55665 | 1627 | apply (rule d(1)) | 
| 53781 | 1628 | apply rule | 
| 1629 | apply rule | |
| 1630 | proof - | |
| 61165 | 1631 | fix y | 
| 1632 | assume "0 < dist y (f x) \<and> dist y (f x) < d" | |
| 53781 | 1633 | then have "g y \<in> g ` f ` (ball x e \<inter> s)" | 
| 55665 | 1634 | using d(2)[unfolded subset_eq,THEN bspec[where x=y]] | 
| 53781 | 1635 | by (auto simp add: dist_commute) | 
| 1636 | then have "g y \<in> ball x e \<inter> s" | |
| 1637 | using assms(4) by auto | |
| 1638 | then show "dist (g y) (g (f x)) < e" | |
| 60420 | 1639 | using assms(4)[rule_format,OF \<open>x \<in> s\<close>] | 
| 44123 | 1640 | by (auto simp add: dist_commute) | 
| 1641 | qed | |
| 1642 | qed | |
| 1643 | moreover have "f x \<in> interior (f ` s)" | |
| 53781 | 1644 | apply (rule sussmann_open_mapping) | 
| 1645 | apply (rule assms ling)+ | |
| 60420 | 1646 | using interior_open[OF assms(1)] and \<open>x \<in> s\<close> | 
| 53781 | 1647 | apply auto | 
| 1648 | done | |
| 61165 | 1649 | moreover have "f (g y) = y" if "y \<in> interior (f ` s)" for y | 
| 53781 | 1650 | proof - | 
| 61165 | 1651 | from that have "y \<in> f ` s" | 
| 53781 | 1652 | using interior_subset by auto | 
| 55665 | 1653 | then obtain z where "z \<in> s" "y = f z" unfolding image_iff .. | 
| 61165 | 1654 | then show ?thesis | 
| 53781 | 1655 | using assms(4) by auto | 
| 44123 | 1656 | qed | 
| 55970 
6d123f0ae358
Some new proofs. Tidying up, esp to remove "apply rule".
 paulson <lp15@cam.ac.uk> parents: 
55665diff
changeset | 1657 | ultimately show ?thesis using assms | 
| 
6d123f0ae358
Some new proofs. Tidying up, esp to remove "apply rule".
 paulson <lp15@cam.ac.uk> parents: 
55665diff
changeset | 1658 | by (metis has_derivative_inverse_basic_x open_interior) | 
| 44123 | 1659 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1660 | |
| 60420 | 1661 | text \<open>A rewrite based on the other domain.\<close> | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1662 | |
| 44123 | 1663 | lemma has_derivative_inverse_strong_x: | 
| 56117 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55970diff
changeset | 1664 | fixes f :: "'a::euclidean_space \<Rightarrow> 'a" | 
| 53781 | 1665 | assumes "open s" | 
| 1666 | and "g y \<in> s" | |
| 1667 | and "continuous_on s f" | |
| 1668 | and "\<forall>x\<in>s. g (f x) = x" | |
| 1669 | and "(f has_derivative f') (at (g y))" | |
| 1670 | and "f' \<circ> g' = id" | |
| 1671 | and "f (g y) = y" | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1672 | shows "(g has_derivative g') (at y)" | 
| 53781 | 1673 | using has_derivative_inverse_strong[OF assms(1-6)] | 
| 1674 | unfolding assms(7) | |
| 1675 | by simp | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1676 | |
| 60420 | 1677 | text \<open>On a region.\<close> | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1678 | |
| 44123 | 1679 | lemma has_derivative_inverse_on: | 
| 56117 
2dbf84ee3deb
remove ordered_euclidean_space constraint from brouwer/derivative lemmas;
 huffman parents: 
55970diff
changeset | 1680 | fixes f :: "'n::euclidean_space \<Rightarrow> 'n" | 
| 53781 | 1681 | assumes "open s" | 
| 1682 | and "\<forall>x\<in>s. (f has_derivative f'(x)) (at x)" | |
| 1683 | and "\<forall>x\<in>s. g (f x) = x" | |
| 1684 | and "f' x \<circ> g' x = id" | |
| 1685 | and "x \<in> s" | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1686 | shows "(g has_derivative g'(x)) (at (f x))" | 
| 53781 | 1687 | apply (rule has_derivative_inverse_strong[where g'="g' x" and f=f]) | 
| 1688 | apply (rule assms)+ | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1689 | unfolding continuous_on_eq_continuous_at[OF assms(1)] | 
| 53781 | 1690 | apply rule | 
| 1691 | apply (rule differentiable_imp_continuous_within) | |
| 1692 | unfolding differentiable_def | |
| 1693 | using assms | |
| 1694 | apply auto | |
| 1695 | done | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1696 | |
| 60420 | 1697 | text \<open>Invertible derivative continous at a point implies local | 
| 44123 | 1698 | injectivity. It's only for this we need continuity of the derivative, | 
| 1699 | except of course if we want the fact that the inverse derivative is | |
| 1700 | also continuous. So if we know for some other reason that the inverse | |
| 60420 | 1701 | function exists, it's OK.\<close> | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1702 | |
| 62381 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62207diff
changeset | 1703 | proposition has_derivative_locally_injective: | 
| 53781 | 1704 | fixes f :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space" | 
| 1705 | assumes "a \<in> s" | |
| 62381 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62207diff
changeset | 1706 | and "open s" | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62207diff
changeset | 1707 | and "bounded_linear g'" | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62207diff
changeset | 1708 | and "g' \<circ> f' a = id" | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62207diff
changeset | 1709 | and "\<And>x. x \<in> s \<Longrightarrow> (f has_derivative f' x) (at x)" | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62207diff
changeset | 1710 | and "\<And>e. e > 0 \<Longrightarrow> \<exists>d>0. \<forall>x. dist a x < d \<longrightarrow> onorm (\<lambda>v. f' x v - f' a v) < e" | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62207diff
changeset | 1711 | obtains r where "r > 0" "ball a r \<subseteq> s" "inj_on f (ball a r)" | 
| 53781 | 1712 | proof - | 
| 1713 | interpret bounded_linear g' | |
| 1714 | using assms by auto | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1715 | note f'g' = assms(4)[unfolded id_def o_def,THEN cong] | 
| 53781 | 1716 | have "g' (f' a (\<Sum>Basis)) = (\<Sum>Basis)" "(\<Sum>Basis) \<noteq> (0::'n)" | 
| 1717 | defer | |
| 1718 | apply (subst euclidean_eq_iff) | |
| 1719 | using f'g' | |
| 1720 | apply auto | |
| 1721 | done | |
| 1722 | then have *: "0 < onorm g'" | |
| 56223 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
56217diff
changeset | 1723 | unfolding onorm_pos_lt[OF assms(3)] | 
| 53781 | 1724 | by fastforce | 
| 1725 | def k \<equiv> "1 / onorm g' / 2" | |
| 1726 | have *: "k > 0" | |
| 1727 | unfolding k_def using * by auto | |
| 55665 | 1728 | obtain d1 where d1: | 
| 1729 | "0 < d1" | |
| 1730 | "\<And>x. dist a x < d1 \<Longrightarrow> onorm (\<lambda>v. f' x v - f' a v) < k" | |
| 1731 | using assms(6) * by blast | |
| 60420 | 1732 | from \<open>open s\<close> obtain d2 where "d2 > 0" "ball a d2 \<subseteq> s" | 
| 1733 | using \<open>a\<in>s\<close> .. | |
| 53781 | 1734 | obtain d2 where "d2 > 0" "ball a d2 \<subseteq> s" | 
| 1735 | using assms(2,1) .. | |
| 55665 | 1736 | obtain d2 where d2: "0 < d2" "ball a d2 \<subseteq> s" | 
| 1737 | using assms(2) | |
| 1738 | unfolding open_contains_ball | |
| 60420 | 1739 | using \<open>a\<in>s\<close> by blast | 
| 55665 | 1740 | obtain d where d: "0 < d" "d < d1" "d < d2" | 
| 1741 | using real_lbound_gt_zero[OF d1(1) d2(1)] by blast | |
| 44123 | 1742 | show ?thesis | 
| 1743 | proof | |
| 62381 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62207diff
changeset | 1744 | show "0 < d" by (fact d) | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62207diff
changeset | 1745 | show "ball a d \<subseteq> s" | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62207diff
changeset | 1746 | using \<open>d < d2\<close> \<open>ball a d2 \<subseteq> s\<close> by auto | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62207diff
changeset | 1747 | show "inj_on f (ball a d)" | 
| 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62207diff
changeset | 1748 | unfolding inj_on_def | 
| 44123 | 1749 | proof (intro strip) | 
| 53781 | 1750 | fix x y | 
| 1751 | assume as: "x \<in> ball a d" "y \<in> ball a d" "f x = f y" | |
| 1752 | def ph \<equiv> "\<lambda>w. w - g' (f w - f x)" | |
| 44123 | 1753 | have ph':"ph = g' \<circ> (\<lambda>w. f' a w - (f w - f x))" | 
| 53781 | 1754 | unfolding ph_def o_def | 
| 1755 | unfolding diff | |
| 1756 | using f'g' | |
| 62381 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62207diff
changeset | 1757 | by (auto simp: algebra_simps) | 
| 53781 | 1758 | have "norm (ph x - ph y) \<le> (1 / 2) * norm (x - y)" | 
| 1759 | apply (rule differentiable_bound[OF convex_ball _ _ as(1-2), where f'="\<lambda>x v. v - g'(f' x v)"]) | |
| 1760 | apply (rule_tac[!] ballI) | |
| 1761 | proof - | |
| 1762 | fix u | |
| 1763 | assume u: "u \<in> ball a d" | |
| 1764 | then have "u \<in> s" | |
| 1765 | using d d2 by auto | |
| 1766 | have *: "(\<lambda>v. v - g' (f' u v)) = g' \<circ> (\<lambda>w. f' a w - f' u w)" | |
| 1767 | unfolding o_def and diff | |
| 1768 | using f'g' by auto | |
| 41958 | 1769 | show "(ph has_derivative (\<lambda>v. v - g' (f' u v))) (at u within ball a d)" | 
| 51642 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1770 | unfolding ph' * | 
| 53781 | 1771 | apply (simp add: comp_def) | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
56151diff
changeset | 1772 | apply (rule bounded_linear.has_derivative[OF assms(3)]) | 
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 1773 | apply (rule derivative_intros) | 
| 53781 | 1774 | defer | 
| 1775 | apply (rule has_derivative_sub[where g'="\<lambda>x.0",unfolded diff_0_right]) | |
| 1776 | apply (rule has_derivative_at_within) | |
| 60420 | 1777 | using assms(5) and \<open>u \<in> s\<close> \<open>a \<in> s\<close> | 
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 1778 | apply (auto intro!: derivative_intros bounded_linear.has_derivative[of _ "\<lambda>x. x"] has_derivative_bounded_linear) | 
| 51642 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1779 | done | 
| 53781 | 1780 | have **: "bounded_linear (\<lambda>x. f' u x - f' a x)" "bounded_linear (\<lambda>x. f' a x - f' u x)" | 
| 1781 | apply (rule_tac[!] bounded_linear_sub) | |
| 56369 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 hoelzl parents: 
56332diff
changeset | 1782 | apply (rule_tac[!] has_derivative_bounded_linear) | 
| 60420 | 1783 | using assms(5) \<open>u \<in> s\<close> \<open>a \<in> s\<close> | 
| 53781 | 1784 | apply auto | 
| 1785 | done | |
| 44123 | 1786 | have "onorm (\<lambda>v. v - g' (f' u v)) \<le> onorm g' * onorm (\<lambda>w. f' a w - f' u w)" | 
| 53781 | 1787 | unfolding * | 
| 1788 | apply (rule onorm_compose) | |
| 1789 | apply (rule assms(3) **)+ | |
| 1790 | done | |
| 44123 | 1791 | also have "\<dots> \<le> onorm g' * k" | 
| 53781 | 1792 | apply (rule mult_left_mono) | 
| 55665 | 1793 | using d1(2)[of u] | 
| 56223 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
56217diff
changeset | 1794 | using onorm_neg[where f="\<lambda>x. f' u x - f' a x"] | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
56217diff
changeset | 1795 | using d and u and onorm_pos_le[OF assms(3)] | 
| 62381 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62207diff
changeset | 1796 | apply (auto simp: algebra_simps) | 
| 53781 | 1797 | done | 
| 1798 | also have "\<dots> \<le> 1 / 2" | |
| 1799 | unfolding k_def by auto | |
| 1800 | finally show "onorm (\<lambda>v. v - g' (f' u v)) \<le> 1 / 2" . | |
| 44123 | 1801 | qed | 
| 1802 | moreover have "norm (ph y - ph x) = norm (y - x)" | |
| 53781 | 1803 | apply (rule arg_cong[where f=norm]) | 
| 1804 | unfolding ph_def | |
| 1805 | using diff | |
| 1806 | unfolding as | |
| 1807 | apply auto | |
| 1808 | done | |
| 1809 | ultimately show "x = y" | |
| 1810 | unfolding norm_minus_commute by auto | |
| 44123 | 1811 | qed | 
| 62381 
a6479cb85944
New and revised material for (multivariate) analysis
 paulson <lp15@cam.ac.uk> parents: 
62207diff
changeset | 1812 | qed | 
| 44123 | 1813 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1814 | |
| 53781 | 1815 | |
| 60420 | 1816 | subsection \<open>Uniformly convergent sequence of derivatives\<close> | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1817 | |
| 44123 | 1818 | lemma has_derivative_sequence_lipschitz_lemma: | 
| 60179 | 1819 | fixes f :: "nat \<Rightarrow> 'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" | 
| 44123 | 1820 | assumes "convex s" | 
| 53781 | 1821 | and "\<forall>n. \<forall>x\<in>s. ((f n) has_derivative (f' n x)) (at x within s)" | 
| 1822 | and "\<forall>n\<ge>N. \<forall>x\<in>s. \<forall>h. norm (f' n x h - g' x h) \<le> e * norm h" | |
| 56271 
61b1e3d88e91
generalized theorems about derivatives of limits of sequences of funtions
 huffman parents: 
56264diff
changeset | 1823 | and "0 \<le> e" | 
| 53781 | 1824 | shows "\<forall>m\<ge>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>y\<in>s. norm ((f m x - f n x) - (f m y - f n y)) \<le> 2 * e * norm (x - y)" | 
| 1825 | proof rule+ | |
| 1826 | fix m n x y | |
| 1827 | assume as: "N \<le> m" "N \<le> n" "x \<in> s" "y \<in> s" | |
| 1828 | show "norm ((f m x - f n x) - (f m y - f n y)) \<le> 2 * e * norm (x - y)" | |
| 1829 | apply (rule differentiable_bound[where f'="\<lambda>x h. f' m x h - f' n x h", OF assms(1) _ _ as(3-4)]) | |
| 1830 | apply (rule_tac[!] ballI) | |
| 1831 | proof - | |
| 1832 | fix x | |
| 1833 | assume "x \<in> s" | |
| 44123 | 1834 | show "((\<lambda>a. f m a - f n a) has_derivative (\<lambda>h. f' m x h - f' n x h)) (at x within s)" | 
| 60420 | 1835 | by (rule derivative_intros assms(2)[rule_format] \<open>x\<in>s\<close>)+ | 
| 56271 
61b1e3d88e91
generalized theorems about derivatives of limits of sequences of funtions
 huffman parents: 
56264diff
changeset | 1836 | show "onorm (\<lambda>h. f' m x h - f' n x h) \<le> 2 * e" | 
| 
61b1e3d88e91
generalized theorems about derivatives of limits of sequences of funtions
 huffman parents: 
56264diff
changeset | 1837 | proof (rule onorm_bound) | 
| 53781 | 1838 | fix h | 
| 44123 | 1839 | have "norm (f' m x h - f' n x h) \<le> norm (f' m x h - g' x h) + norm (f' n x h - g' x h)" | 
| 1840 | using norm_triangle_ineq[of "f' m x h - g' x h" "- f' n x h + g' x h"] | |
| 53781 | 1841 | unfolding norm_minus_commute | 
| 1842 | by (auto simp add: algebra_simps) | |
| 1843 | also have "\<dots> \<le> e * norm h + e * norm h" | |
| 60420 | 1844 | using assms(3)[rule_format,OF \<open>N \<le> m\<close> \<open>x \<in> s\<close>, of h] | 
| 1845 | using assms(3)[rule_format,OF \<open>N \<le> n\<close> \<open>x \<in> s\<close>, of h] | |
| 53781 | 1846 | by (auto simp add: field_simps) | 
| 56271 
61b1e3d88e91
generalized theorems about derivatives of limits of sequences of funtions
 huffman parents: 
56264diff
changeset | 1847 | finally show "norm (f' m x h - f' n x h) \<le> 2 * e * norm h" | 
| 53781 | 1848 | by auto | 
| 60420 | 1849 | qed (simp add: \<open>0 \<le> e\<close>) | 
| 44123 | 1850 | qed | 
| 1851 | qed | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1852 | |
| 44123 | 1853 | lemma has_derivative_sequence_lipschitz: | 
| 60179 | 1854 | fixes f :: "nat \<Rightarrow> 'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" | 
| 44123 | 1855 | assumes "convex s" | 
| 53781 | 1856 | and "\<forall>n. \<forall>x\<in>s. ((f n) has_derivative (f' n x)) (at x within s)" | 
| 1857 | and "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>h. norm (f' n x h - g' x h) \<le> e * norm h" | |
| 1858 | shows "\<forall>e>0. \<exists>N. \<forall>m\<ge>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>y\<in>s. | |
| 1859 | norm ((f m x - f n x) - (f m y - f n y)) \<le> e * norm (x - y)" | |
| 1860 | proof (rule, rule) | |
| 61165 | 1861 | fix e :: real | 
| 1862 | assume "e > 0" | |
| 1863 | then have *: "2 * (1/2* e) = e" "1/2 * e >0" | |
| 1864 | by auto | |
| 55665 | 1865 | obtain N where "\<forall>n\<ge>N. \<forall>x\<in>s. \<forall>h. norm (f' n x h - g' x h) \<le> 1 / 2 * e * norm h" | 
| 1866 | using assms(3) *(2) by blast | |
| 61165 | 1867 | then show "\<exists>N. \<forall>m\<ge>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>y\<in>s. norm (f m x - f n x - (f m y - f n y)) \<le> e * norm (x - y)" | 
| 53781 | 1868 | apply (rule_tac x=N in exI) | 
| 1869 | apply (rule has_derivative_sequence_lipschitz_lemma[where e="1/2 *e", unfolded *]) | |
| 60420 | 1870 | using assms \<open>e > 0\<close> | 
| 53781 | 1871 | apply auto | 
| 1872 | done | |
| 44123 | 1873 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1874 | |
| 44123 | 1875 | lemma has_derivative_sequence: | 
| 60179 | 1876 | fixes f :: "nat \<Rightarrow> 'a::real_normed_vector \<Rightarrow> 'b::banach" | 
| 44123 | 1877 | assumes "convex s" | 
| 53781 | 1878 | and "\<forall>n. \<forall>x\<in>s. ((f n) has_derivative (f' n x)) (at x within s)" | 
| 1879 | and "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>h. norm (f' n x h - g' x h) \<le> e * norm h" | |
| 1880 | and "x0 \<in> s" | |
| 61973 | 1881 | and "((\<lambda>n. f n x0) \<longlongrightarrow> l) sequentially" | 
| 1882 | shows "\<exists>g. \<forall>x\<in>s. ((\<lambda>n. f n x) \<longlongrightarrow> g x) sequentially \<and> (g has_derivative g'(x)) (at x within s)" | |
| 53781 | 1883 | proof - | 
| 1884 | have lem1: "\<forall>e>0. \<exists>N. \<forall>m\<ge>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>y\<in>s. | |
| 1885 | norm ((f m x - f n x) - (f m y - f n y)) \<le> e * norm (x - y)" | |
| 56320 | 1886 | using assms(1,2,3) by (rule has_derivative_sequence_lipschitz) | 
| 61973 | 1887 | have "\<exists>g. \<forall>x\<in>s. ((\<lambda>n. f n x) \<longlongrightarrow> g x) sequentially" | 
| 53781 | 1888 | apply (rule bchoice) | 
| 1889 | unfolding convergent_eq_cauchy | |
| 44123 | 1890 | proof | 
| 53781 | 1891 | fix x | 
| 1892 | assume "x \<in> s" | |
| 1893 | show "Cauchy (\<lambda>n. f n x)" | |
| 1894 | proof (cases "x = x0") | |
| 1895 | case True | |
| 1896 | then show ?thesis | |
| 1897 | using LIMSEQ_imp_Cauchy[OF assms(5)] by auto | |
| 44123 | 1898 | next | 
| 53781 | 1899 | case False | 
| 1900 | show ?thesis | |
| 1901 | unfolding Cauchy_def | |
| 1902 | proof (rule, rule) | |
| 1903 | fix e :: real | |
| 1904 | assume "e > 0" | |
| 56541 | 1905 | hence *: "e / 2 > 0" "e / 2 / norm (x - x0) > 0" using False by auto | 
| 55665 | 1906 | obtain M where M: "\<forall>m\<ge>M. \<forall>n\<ge>M. dist (f m x0) (f n x0) < e / 2" | 
| 1907 | using LIMSEQ_imp_Cauchy[OF assms(5)] | |
| 1908 | unfolding Cauchy_def | |
| 1909 | using *(1) by blast | |
| 1910 | obtain N where N: | |
| 1911 | "\<forall>m\<ge>N. \<forall>n\<ge>N. | |
| 1912 | \<forall>xa\<in>s. \<forall>y\<in>s. norm (f m xa - f n xa - (f m y - f n y)) \<le> | |
| 1913 | e / 2 / norm (x - x0) * norm (xa - y)" | |
| 1914 | using lem1 *(2) by blast | |
| 44123 | 1915 | show "\<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. dist (f m x) (f n x) < e" | 
| 53781 | 1916 | apply (rule_tac x="max M N" in exI) | 
| 1917 | proof rule+ | |
| 1918 | fix m n | |
| 1919 | assume as: "max M N \<le>m" "max M N\<le>n" | |
| 1920 | have "dist (f m x) (f n x) \<le> | |
| 1921 | norm (f m x0 - f n x0) + norm (f m x - f n x - (f m x0 - f n x0))" | |
| 1922 | unfolding dist_norm | |
| 1923 | by (rule norm_triangle_sub) | |
| 44123 | 1924 | also have "\<dots> \<le> norm (f m x0 - f n x0) + e / 2" | 
| 60420 | 1925 | using N[rule_format,OF _ _ \<open>x\<in>s\<close> \<open>x0\<in>s\<close>, of m n] and as and False | 
| 44123 | 1926 | by auto | 
| 1927 | also have "\<dots> < e / 2 + e / 2" | |
| 53781 | 1928 | apply (rule add_strict_right_mono) | 
| 1929 | using as and M[rule_format] | |
| 1930 | unfolding dist_norm | |
| 1931 | apply auto | |
| 1932 | done | |
| 1933 | finally show "dist (f m x) (f n x) < e" | |
| 1934 | by auto | |
| 44123 | 1935 | qed | 
| 1936 | qed | |
| 1937 | qed | |
| 1938 | qed | |
| 61969 | 1939 | then obtain g where g: "\<forall>x\<in>s. (\<lambda>n. f n x) \<longlonglongrightarrow> g x" .. | 
| 53781 | 1940 | have lem2: "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>y\<in>s. norm ((f n x - f n y) - (g x - g y)) \<le> e * norm (x - y)" | 
| 1941 | proof (rule, rule) | |
| 1942 | fix e :: real | |
| 1943 | assume *: "e > 0" | |
| 55665 | 1944 | obtain N where | 
| 1945 | N: "\<forall>m\<ge>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>y\<in>s. norm (f m x - f n x - (f m y - f n y)) \<le> e * norm (x - y)" | |
| 1946 | using lem1 * by blast | |
| 44123 | 1947 | show "\<exists>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>y\<in>s. norm (f n x - f n y - (g x - g y)) \<le> e * norm (x - y)" | 
| 53781 | 1948 | apply (rule_tac x=N in exI) | 
| 1949 | proof rule+ | |
| 1950 | fix n x y | |
| 1951 | assume as: "N \<le> n" "x \<in> s" "y \<in> s" | |
| 61973 | 1952 | have "((\<lambda>m. norm (f n x - f n y - (f m x - f m y))) \<longlongrightarrow> norm (f n x - f n y - (g x - g y))) sequentially" | 
| 56320 | 1953 | by (intro tendsto_intros g[rule_format] as) | 
| 1954 | moreover have "eventually (\<lambda>m. norm (f n x - f n y - (f m x - f m y)) \<le> e * norm (x - y)) sequentially" | |
| 44123 | 1955 | unfolding eventually_sequentially | 
| 53781 | 1956 | apply (rule_tac x=N in exI) | 
| 1957 | apply rule | |
| 1958 | apply rule | |
| 1959 | proof - | |
| 1960 | fix m | |
| 1961 | assume "N \<le> m" | |
| 1962 | then show "norm (f n x - f n y - (f m x - f m y)) \<le> e * norm (x - y)" | |
| 44123 | 1963 | using N[rule_format, of n m x y] and as | 
| 1964 | by (auto simp add: algebra_simps) | |
| 1965 | qed | |
| 56320 | 1966 | ultimately show "norm (f n x - f n y - (g x - g y)) \<le> e * norm (x - y)" | 
| 1967 | by (rule tendsto_ge_const[OF trivial_limit_sequentially]) | |
| 44123 | 1968 | qed | 
| 1969 | qed | |
| 61973 | 1970 | have "\<forall>x\<in>s. ((\<lambda>n. f n x) \<longlongrightarrow> g x) sequentially \<and> (g has_derivative g' x) (at x within s)" | 
| 56320 | 1971 | unfolding has_derivative_within_alt2 | 
| 1972 | proof (intro ballI conjI) | |
| 53781 | 1973 | fix x | 
| 1974 | assume "x \<in> s" | |
| 61973 | 1975 | then show "((\<lambda>n. f n x) \<longlongrightarrow> g x) sequentially" | 
| 56320 | 1976 | by (simp add: g) | 
| 61973 | 1977 | have lem3: "\<forall>u. ((\<lambda>n. f' n x u) \<longlongrightarrow> g' x u) sequentially" | 
| 56320 | 1978 | unfolding filterlim_def le_nhds_metric_le eventually_filtermap dist_norm | 
| 1979 | proof (intro allI impI) | |
| 53781 | 1980 | fix u | 
| 1981 | fix e :: real | |
| 1982 | assume "e > 0" | |
| 56320 | 1983 | show "eventually (\<lambda>n. norm (f' n x u - g' x u) \<le> e) sequentially" | 
| 53781 | 1984 | proof (cases "u = 0") | 
| 1985 | case True | |
| 56320 | 1986 | have "eventually (\<lambda>n. norm (f' n x u - g' x u) \<le> e * norm u) sequentially" | 
| 60420 | 1987 | using assms(3)[folded eventually_sequentially] and \<open>0 < e\<close> and \<open>x \<in> s\<close> | 
| 61810 | 1988 | by (fast elim: eventually_mono) | 
| 56320 | 1989 | then show ?thesis | 
| 61810 | 1990 | using \<open>u = 0\<close> and \<open>0 < e\<close> by (auto elim: eventually_mono) | 
| 44123 | 1991 | next | 
| 53781 | 1992 | case False | 
| 60420 | 1993 | with \<open>0 < e\<close> have "0 < e / norm u" by simp | 
| 56320 | 1994 | then have "eventually (\<lambda>n. norm (f' n x u - g' x u) \<le> e / norm u * norm u) sequentially" | 
| 60420 | 1995 | using assms(3)[folded eventually_sequentially] and \<open>x \<in> s\<close> | 
| 61810 | 1996 | by (fast elim: eventually_mono) | 
| 56320 | 1997 | then show ?thesis | 
| 60420 | 1998 | using \<open>u \<noteq> 0\<close> by simp | 
| 44123 | 1999 | qed | 
| 2000 | qed | |
| 2001 | show "bounded_linear (g' x)" | |
| 56271 
61b1e3d88e91
generalized theorems about derivatives of limits of sequences of funtions
 huffman parents: 
56264diff
changeset | 2002 | proof | 
| 
61b1e3d88e91
generalized theorems about derivatives of limits of sequences of funtions
 huffman parents: 
56264diff
changeset | 2003 | fix x' y z :: 'a | 
| 53781 | 2004 | fix c :: real | 
| 60420 | 2005 | note lin = assms(2)[rule_format,OF \<open>x\<in>s\<close>,THEN has_derivative_bounded_linear] | 
| 44123 | 2006 | show "g' x (c *\<^sub>R x') = c *\<^sub>R g' x x'" | 
| 53781 | 2007 | apply (rule tendsto_unique[OF trivial_limit_sequentially]) | 
| 2008 | apply (rule lem3[rule_format]) | |
| 56369 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 hoelzl parents: 
56332diff
changeset | 2009 | unfolding lin[THEN bounded_linear.linear, THEN linear_cmul] | 
| 53781 | 2010 | apply (intro tendsto_intros) | 
| 2011 | apply (rule lem3[rule_format]) | |
| 2012 | done | |
| 44123 | 2013 | show "g' x (y + z) = g' x y + g' x z" | 
| 53781 | 2014 | apply (rule tendsto_unique[OF trivial_limit_sequentially]) | 
| 2015 | apply (rule lem3[rule_format]) | |
| 56369 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 hoelzl parents: 
56332diff
changeset | 2016 | unfolding lin[THEN bounded_linear.linear, THEN linear_add] | 
| 53781 | 2017 | apply (rule tendsto_add) | 
| 2018 | apply (rule lem3[rule_format])+ | |
| 2019 | done | |
| 56271 
61b1e3d88e91
generalized theorems about derivatives of limits of sequences of funtions
 huffman parents: 
56264diff
changeset | 2020 | obtain N where N: "\<forall>h. norm (f' N x h - g' x h) \<le> 1 * norm h" | 
| 60420 | 2021 | using assms(3) \<open>x \<in> s\<close> by (fast intro: zero_less_one) | 
| 56271 
61b1e3d88e91
generalized theorems about derivatives of limits of sequences of funtions
 huffman parents: 
56264diff
changeset | 2022 | have "bounded_linear (f' N x)" | 
| 60420 | 2023 | using assms(2) \<open>x \<in> s\<close> by fast | 
| 56271 
61b1e3d88e91
generalized theorems about derivatives of limits of sequences of funtions
 huffman parents: 
56264diff
changeset | 2024 | from bounded_linear.bounded [OF this] | 
| 
61b1e3d88e91
generalized theorems about derivatives of limits of sequences of funtions
 huffman parents: 
56264diff
changeset | 2025 | obtain K where K: "\<forall>h. norm (f' N x h) \<le> norm h * K" .. | 
| 
61b1e3d88e91
generalized theorems about derivatives of limits of sequences of funtions
 huffman parents: 
56264diff
changeset | 2026 |       {
 | 
| 
61b1e3d88e91
generalized theorems about derivatives of limits of sequences of funtions
 huffman parents: 
56264diff
changeset | 2027 | fix h | 
| 
61b1e3d88e91
generalized theorems about derivatives of limits of sequences of funtions
 huffman parents: 
56264diff
changeset | 2028 | have "norm (g' x h) = norm (f' N x h - (f' N x h - g' x h))" | 
| 
61b1e3d88e91
generalized theorems about derivatives of limits of sequences of funtions
 huffman parents: 
56264diff
changeset | 2029 | by simp | 
| 
61b1e3d88e91
generalized theorems about derivatives of limits of sequences of funtions
 huffman parents: 
56264diff
changeset | 2030 | also have "\<dots> \<le> norm (f' N x h) + norm (f' N x h - g' x h)" | 
| 
61b1e3d88e91
generalized theorems about derivatives of limits of sequences of funtions
 huffman parents: 
56264diff
changeset | 2031 | by (rule norm_triangle_ineq4) | 
| 
61b1e3d88e91
generalized theorems about derivatives of limits of sequences of funtions
 huffman parents: 
56264diff
changeset | 2032 | also have "\<dots> \<le> norm h * K + 1 * norm h" | 
| 
61b1e3d88e91
generalized theorems about derivatives of limits of sequences of funtions
 huffman parents: 
56264diff
changeset | 2033 | using N K by (fast intro: add_mono) | 
| 
61b1e3d88e91
generalized theorems about derivatives of limits of sequences of funtions
 huffman parents: 
56264diff
changeset | 2034 | finally have "norm (g' x h) \<le> norm h * (K + 1)" | 
| 
61b1e3d88e91
generalized theorems about derivatives of limits of sequences of funtions
 huffman parents: 
56264diff
changeset | 2035 | by (simp add: ring_distribs) | 
| 
61b1e3d88e91
generalized theorems about derivatives of limits of sequences of funtions
 huffman parents: 
56264diff
changeset | 2036 | } | 
| 
61b1e3d88e91
generalized theorems about derivatives of limits of sequences of funtions
 huffman parents: 
56264diff
changeset | 2037 | then show "\<exists>K. \<forall>h. norm (g' x h) \<le> norm h * K" by fast | 
| 44123 | 2038 | qed | 
| 56320 | 2039 | show "\<forall>e>0. eventually (\<lambda>y. norm (g y - g x - g' x (y - x)) \<le> e * norm (y - x)) (at x within s)" | 
| 53781 | 2040 | proof (rule, rule) | 
| 61165 | 2041 | fix e :: real | 
| 2042 | assume "e > 0" | |
| 2043 | then have *: "e / 3 > 0" | |
| 2044 | by auto | |
| 55665 | 2045 | obtain N1 where N1: "\<forall>n\<ge>N1. \<forall>x\<in>s. \<forall>h. norm (f' n x h - g' x h) \<le> e / 3 * norm h" | 
| 2046 | using assms(3) * by blast | |
| 2047 | obtain N2 where | |
| 2048 | N2: "\<forall>n\<ge>N2. \<forall>x\<in>s. \<forall>y\<in>s. norm (f n x - f n y - (g x - g y)) \<le> e / 3 * norm (x - y)" | |
| 2049 | using lem2 * by blast | |
| 56320 | 2050 | let ?N = "max N1 N2" | 
| 2051 | have "eventually (\<lambda>y. norm (f ?N y - f ?N x - f' ?N x (y - x)) \<le> e / 3 * norm (y - x)) (at x within s)" | |
| 60420 | 2052 | using assms(2)[unfolded has_derivative_within_alt2] and \<open>x \<in> s\<close> and * by fast | 
| 56320 | 2053 | moreover have "eventually (\<lambda>y. y \<in> s) (at x within s)" | 
| 2054 | unfolding eventually_at by (fast intro: zero_less_one) | |
| 61165 | 2055 | ultimately show "\<forall>\<^sub>F y in at x within s. norm (g y - g x - g' x (y - x)) \<le> e * norm (y - x)" | 
| 56320 | 2056 | proof (rule eventually_elim2) | 
| 53781 | 2057 | fix y | 
| 56320 | 2058 | assume "y \<in> s" | 
| 2059 | assume "norm (f ?N y - f ?N x - f' ?N x (y - x)) \<le> e / 3 * norm (y - x)" | |
| 2060 | moreover have "norm (g y - g x - (f ?N y - f ?N x)) \<le> e / 3 * norm (y - x)" | |
| 60420 | 2061 | using N2[rule_format, OF _ \<open>y \<in> s\<close> \<open>x \<in> s\<close>] | 
| 56320 | 2062 | by (simp add: norm_minus_commute) | 
| 2063 | ultimately have "norm (g y - g x - f' ?N x (y - x)) \<le> 2 * e / 3 * norm (y - x)" | |
| 44123 | 2064 | using norm_triangle_le[of "g y - g x - (f ?N y - f ?N x)" "f ?N y - f ?N x - f' ?N x (y - x)" "2 * e / 3 * norm (y - x)"] | 
| 53781 | 2065 | by (auto simp add: algebra_simps) | 
| 44123 | 2066 | moreover | 
| 2067 | have " norm (f' ?N x (y - x) - g' x (y - x)) \<le> e / 3 * norm (y - x)" | |
| 60420 | 2068 | using N1 \<open>x \<in> s\<close> by auto | 
| 41958 | 2069 | ultimately show "norm (g y - g x - g' x (y - x)) \<le> e * norm (y - x)" | 
| 44123 | 2070 | using norm_triangle_le[of "g y - g x - f' (max N1 N2) x (y - x)" "f' (max N1 N2) x (y - x) - g' x (y - x)"] | 
| 53781 | 2071 | by (auto simp add: algebra_simps) | 
| 44123 | 2072 | qed | 
| 2073 | qed | |
| 2074 | qed | |
| 56320 | 2075 | then show ?thesis by fast | 
| 44123 | 2076 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2077 | |
| 60420 | 2078 | text \<open>Can choose to line up antiderivatives if we want.\<close> | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2079 | |
| 44123 | 2080 | lemma has_antiderivative_sequence: | 
| 60179 | 2081 | fixes f :: "nat \<Rightarrow> 'a::real_normed_vector \<Rightarrow> 'b::banach" | 
| 44123 | 2082 | assumes "convex s" | 
| 53781 | 2083 | and "\<forall>n. \<forall>x\<in>s. ((f n) has_derivative (f' n x)) (at x within s)" | 
| 2084 | and "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>h. norm (f' n x h - g' x h) \<le> e * norm h" | |
| 2085 | shows "\<exists>g. \<forall>x\<in>s. (g has_derivative g' x) (at x within s)" | |
| 2086 | proof (cases "s = {}")
 | |
| 2087 | case False | |
| 2088 | then obtain a where "a \<in> s" | |
| 2089 | by auto | |
| 2090 | have *: "\<And>P Q. \<exists>g. \<forall>x\<in>s. P g x \<and> Q g x \<Longrightarrow> \<exists>g. \<forall>x\<in>s. Q g x" | |
| 2091 | by auto | |
| 44123 | 2092 | show ?thesis | 
| 53781 | 2093 | apply (rule *) | 
| 2094 | apply (rule has_derivative_sequence[OF assms(1) _ assms(3), of "\<lambda>n x. f n x + (f 0 a - f n a)"]) | |
| 55970 
6d123f0ae358
Some new proofs. Tidying up, esp to remove "apply rule".
 paulson <lp15@cam.ac.uk> parents: 
55665diff
changeset | 2095 | apply (metis assms(2) has_derivative_add_const) | 
| 60420 | 2096 | apply (rule \<open>a \<in> s\<close>) | 
| 53781 | 2097 | apply auto | 
| 2098 | done | |
| 44123 | 2099 | qed auto | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2100 | |
| 44123 | 2101 | lemma has_antiderivative_limit: | 
| 60179 | 2102 | fixes g' :: "'a::real_normed_vector \<Rightarrow> 'a \<Rightarrow> 'b::banach" | 
| 44123 | 2103 | assumes "convex s" | 
| 53781 | 2104 | and "\<forall>e>0. \<exists>f f'. \<forall>x\<in>s. | 
| 2105 | (f has_derivative (f' x)) (at x within s) \<and> (\<forall>h. norm (f' x h - g' x h) \<le> e * norm h)" | |
| 2106 | shows "\<exists>g. \<forall>x\<in>s. (g has_derivative g' x) (at x within s)" | |
| 2107 | proof - | |
| 2108 | have *: "\<forall>n. \<exists>f f'. \<forall>x\<in>s. | |
| 2109 | (f has_derivative (f' x)) (at x within s) \<and> | |
| 2110 | (\<forall>h. norm(f' x h - g' x h) \<le> inverse (real (Suc n)) * norm h)" | |
| 61649 
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
 paulson <lp15@cam.ac.uk> parents: 
61560diff
changeset | 2111 | by (simp add: assms(2)) | 
| 55665 | 2112 | obtain f where | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
56151diff
changeset | 2113 | *: "\<forall>x. \<exists>f'. \<forall>xa\<in>s. (f x has_derivative f' xa) (at xa within s) \<and> | 
| 55665 | 2114 | (\<forall>h. norm (f' xa h - g' xa h) \<le> inverse (real (Suc x)) * norm h)" | 
| 2115 | using *[THEN choice] .. | |
| 2116 | obtain f' where | |
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
56151diff
changeset | 2117 | f: "\<forall>x. \<forall>xa\<in>s. (f x has_derivative f' x xa) (at xa within s) \<and> | 
| 55665 | 2118 | (\<forall>h. norm (f' x xa h - g' xa h) \<le> inverse (real (Suc x)) * norm h)" | 
| 2119 | using *[THEN choice] .. | |
| 53781 | 2120 | show ?thesis | 
| 2121 | apply (rule has_antiderivative_sequence[OF assms(1), of f f']) | |
| 2122 | defer | |
| 2123 | apply rule | |
| 2124 | apply rule | |
| 2125 | proof - | |
| 2126 | fix e :: real | |
| 2127 | assume "e > 0" | |
| 55665 | 2128 | obtain N where N: "inverse (real (Suc N)) < e" | 
| 60420 | 2129 | using reals_Archimedean[OF \<open>e>0\<close>] .. | 
| 44123 | 2130 | show "\<exists>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>h. norm (f' n x h - g' x h) \<le> e * norm h" | 
| 53781 | 2131 | apply (rule_tac x=N in exI) | 
| 61165 | 2132 | apply rule | 
| 2133 | apply rule | |
| 2134 | apply rule | |
| 2135 | apply rule | |
| 2136 | proof - | |
| 2137 | fix n x h | |
| 2138 | assume n: "N \<le> n" and x: "x \<in> s" | |
| 53781 | 2139 | have *: "inverse (real (Suc n)) \<le> e" | 
| 2140 | apply (rule order_trans[OF _ N[THEN less_imp_le]]) | |
| 61165 | 2141 | using n | 
| 53781 | 2142 | apply (auto simp add: field_simps) | 
| 2143 | done | |
| 61165 | 2144 | show "norm (f' n x h - g' x h) \<le> e * norm h" | 
| 2145 | using f[rule_format,THEN conjunct2, OF x, of n, THEN spec[where x=h]] | |
| 53781 | 2146 | apply (rule order_trans) | 
| 2147 | using N * | |
| 2148 | apply (cases "h = 0") | |
| 2149 | apply auto | |
| 2150 | done | |
| 44123 | 2151 | qed | 
| 53781 | 2152 | qed (insert f, auto) | 
| 44123 | 2153 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2154 | |
| 53781 | 2155 | |
| 60420 | 2156 | subsection \<open>Differentiation of a series\<close> | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2157 | |
| 44123 | 2158 | lemma has_derivative_series: | 
| 60179 | 2159 | fixes f :: "nat \<Rightarrow> 'a::real_normed_vector \<Rightarrow> 'b::banach" | 
| 44123 | 2160 | assumes "convex s" | 
| 56183 | 2161 | and "\<And>n x. x \<in> s \<Longrightarrow> ((f n) has_derivative (f' n x)) (at x within s)" | 
| 56193 
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
 hoelzl parents: 
56188diff
changeset | 2162 |     and "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>h. norm (setsum (\<lambda>i. f' i x h) {..<n} - g' x h) \<le> e * norm h"
 | 
| 53781 | 2163 | and "x \<in> s" | 
| 56183 | 2164 | and "(\<lambda>n. f n x) sums l" | 
| 2165 | shows "\<exists>g. \<forall>x\<in>s. (\<lambda>n. f n x) sums (g x) \<and> (g has_derivative g' x) (at x within s)" | |
| 2166 | unfolding sums_def | |
| 53781 | 2167 | apply (rule has_derivative_sequence[OF assms(1) _ assms(3)]) | 
| 55970 
6d123f0ae358
Some new proofs. Tidying up, esp to remove "apply rule".
 paulson <lp15@cam.ac.uk> parents: 
55665diff
changeset | 2168 | apply (metis assms(2) has_derivative_setsum) | 
| 53781 | 2169 | using assms(4-5) | 
| 56183 | 2170 | unfolding sums_def | 
| 53781 | 2171 | apply auto | 
| 2172 | done | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2173 | |
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2174 | lemma has_field_derivative_series: | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2175 |   fixes f :: "nat \<Rightarrow> ('a :: {real_normed_field,banach}) \<Rightarrow> 'a"
 | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2176 | assumes "convex s" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2177 | assumes "\<And>n x. x \<in> s \<Longrightarrow> (f n has_field_derivative f' n x) (at x within s)" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2178 | assumes "uniform_limit s (\<lambda>n x. \<Sum>i<n. f' i x) g' sequentially" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2179 | assumes "x0 \<in> s" "summable (\<lambda>n. f n x0)" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2180 | shows "\<exists>g. \<forall>x\<in>s. (\<lambda>n. f n x) sums g x \<and> (g has_field_derivative g' x) (at x within s)" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2181 | unfolding has_field_derivative_def | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2182 | proof (rule has_derivative_series) | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2183 | show "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>h. norm ((\<Sum>i<n. f' i x * h) - g' x * h) \<le> e * norm h" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2184 | proof (intro allI impI) | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2185 | fix e :: real assume "e > 0" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2186 | with assms(3) obtain N where N: "\<And>n x. n \<ge> N \<Longrightarrow> x \<in> s \<Longrightarrow> norm ((\<Sum>i<n. f' i x) - g' x) < e" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2187 | unfolding uniform_limit_iff eventually_at_top_linorder dist_norm by blast | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2188 |     {
 | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2189 | fix n :: nat and x h :: 'a assume nx: "n \<ge> N" "x \<in> s" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2190 | have "norm ((\<Sum>i<n. f' i x * h) - g' x * h) = norm ((\<Sum>i<n. f' i x) - g' x) * norm h" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2191 | by (simp add: norm_mult [symmetric] ring_distribs setsum_left_distrib) | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2192 | also from N[OF nx] have "norm ((\<Sum>i<n. f' i x) - g' x) \<le> e" by simp | 
| 61649 
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
 paulson <lp15@cam.ac.uk> parents: 
61560diff
changeset | 2193 | hence "norm ((\<Sum>i<n. f' i x) - g' x) * norm h \<le> e * norm h" | 
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2194 | by (intro mult_right_mono) simp_all | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2195 | finally have "norm ((\<Sum>i<n. f' i x * h) - g' x * h) \<le> e * norm h" . | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2196 | } | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2197 | thus "\<exists>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>h. norm ((\<Sum>i<n. f' i x * h) - g' x * h) \<le> e * norm h" by blast | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2198 | qed | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2199 | qed (insert assms, auto simp: has_field_derivative_def) | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2200 | |
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2201 | lemma has_field_derivative_series': | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2202 |   fixes f :: "nat \<Rightarrow> ('a :: {real_normed_field,banach}) \<Rightarrow> 'a"
 | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2203 | assumes "convex s" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2204 | assumes "\<And>n x. x \<in> s \<Longrightarrow> (f n has_field_derivative f' n x) (at x within s)" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2205 | assumes "uniformly_convergent_on s (\<lambda>n x. \<Sum>i<n. f' i x)" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2206 | assumes "x0 \<in> s" "summable (\<lambda>n. f n x0)" "x \<in> interior s" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2207 | shows "summable (\<lambda>n. f n x)" "((\<lambda>x. \<Sum>n. f n x) has_field_derivative (\<Sum>n. f' n x)) (at x)" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2208 | proof - | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2209 | from \<open>x \<in> interior s\<close> have "x \<in> s" using interior_subset by blast | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2210 | def g' \<equiv> "\<lambda>x. \<Sum>i. f' i x" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2211 | from assms(3) have "uniform_limit s (\<lambda>n x. \<Sum>i<n. f' i x) g' sequentially" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2212 | by (simp add: uniformly_convergent_uniform_limit_iff suminf_eq_lim g'_def) | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2213 | from has_field_derivative_series[OF assms(1,2) this assms(4,5)] obtain g where g: | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2214 | "\<And>x. x \<in> s \<Longrightarrow> (\<lambda>n. f n x) sums g x" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2215 | "\<And>x. x \<in> s \<Longrightarrow> (g has_field_derivative g' x) (at x within s)" by blast | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2216 | from g(1)[OF \<open>x \<in> s\<close>] show "summable (\<lambda>n. f n x)" by (simp add: sums_iff) | 
| 61649 
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
 paulson <lp15@cam.ac.uk> parents: 
61560diff
changeset | 2217 | from g(2)[OF \<open>x \<in> s\<close>] \<open>x \<in> interior s\<close> have "(g has_field_derivative g' x) (at x)" | 
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2218 | by (simp add: at_within_interior[of x s]) | 
| 61649 
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
 paulson <lp15@cam.ac.uk> parents: 
61560diff
changeset | 2219 | also have "(g has_field_derivative g' x) (at x) \<longleftrightarrow> | 
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2220 | ((\<lambda>x. \<Sum>n. f n x) has_field_derivative g' x) (at x)" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2221 | using eventually_nhds_in_nhd[OF \<open>x \<in> interior s\<close>] interior_subset[of s] g(1) | 
| 61810 | 2222 | by (intro DERIV_cong_ev) (auto elim!: eventually_mono simp: sums_iff) | 
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2223 | finally show "((\<lambda>x. \<Sum>n. f n x) has_field_derivative g' x) (at x)" . | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2224 | qed | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2225 | |
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2226 | lemma differentiable_series: | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2227 |   fixes f :: "nat \<Rightarrow> ('a :: {real_normed_field,banach}) \<Rightarrow> 'a"
 | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2228 | assumes "convex s" "open s" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2229 | assumes "\<And>n x. x \<in> s \<Longrightarrow> (f n has_field_derivative f' n x) (at x)" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2230 | assumes "uniformly_convergent_on s (\<lambda>n x. \<Sum>i<n. f' i x)" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2231 | assumes "x0 \<in> s" "summable (\<lambda>n. f n x0)" and x: "x \<in> s" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2232 | shows "summable (\<lambda>n. f n x)" and "(\<lambda>x. \<Sum>n. f n x) differentiable (at x)" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2233 | proof - | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2234 | from assms(4) obtain g' where A: "uniform_limit s (\<lambda>n x. \<Sum>i<n. f' i x) g' sequentially" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2235 | unfolding uniformly_convergent_on_def by blast | 
| 61808 | 2236 | from x and \<open>open s\<close> have s: "at x within s = at x" by (rule at_within_open) | 
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2237 | have "\<exists>g. \<forall>x\<in>s. (\<lambda>n. f n x) sums g x \<and> (g has_field_derivative g' x) (at x within s)" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2238 | by (intro has_field_derivative_series[of s f f' g' x0] assms A has_field_derivative_at_within) | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2239 | then obtain g where g: "\<And>x. x \<in> s \<Longrightarrow> (\<lambda>n. f n x) sums g x" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2240 | "\<And>x. x \<in> s \<Longrightarrow> (g has_field_derivative g' x) (at x within s)" by blast | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2241 | from g[OF x] show "summable (\<lambda>n. f n x)" by (auto simp: summable_def) | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2242 | from g(2)[OF x] have g': "(g has_derivative op * (g' x)) (at x)" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2243 | by (simp add: has_field_derivative_def s) | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2244 | have "((\<lambda>x. \<Sum>n. f n x) has_derivative op * (g' x)) (at x)" | 
| 62087 
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
 paulson parents: 
61975diff
changeset | 2245 | by (rule has_derivative_transform_within_open[OF g' \<open>open s\<close> x]) | 
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2246 | (insert g, auto simp: sums_iff) | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2247 | thus "(\<lambda>x. \<Sum>n. f n x) differentiable (at x)" unfolding differentiable_def | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2248 | by (auto simp: summable_def differentiable_def has_field_derivative_def) | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2249 | qed | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2250 | |
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2251 | lemma differentiable_series': | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2252 |   fixes f :: "nat \<Rightarrow> ('a :: {real_normed_field,banach}) \<Rightarrow> 'a"
 | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2253 | assumes "convex s" "open s" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2254 | assumes "\<And>n x. x \<in> s \<Longrightarrow> (f n has_field_derivative f' n x) (at x)" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2255 | assumes "uniformly_convergent_on s (\<lambda>n x. \<Sum>i<n. f' i x)" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2256 | assumes "x0 \<in> s" "summable (\<lambda>n. f n x0)" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2257 | shows "(\<lambda>x. \<Sum>n. f n x) differentiable (at x0)" | 
| 61808 | 2258 | using differentiable_series[OF assms, of x0] \<open>x0 \<in> s\<close> by blast+ | 
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2259 | |
| 61076 | 2260 | text \<open>Considering derivative @{typ "real \<Rightarrow> 'b::real_normed_vector"} as a vector.\<close>
 | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2261 | |
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 2262 | definition "vector_derivative f net = (SOME f'. (f has_vector_derivative f') net)" | 
| 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 2263 | |
| 61245 | 2264 | lemma vector_derivative_unique_within: | 
| 2265 | assumes not_bot: "at x within s \<noteq> bot" | |
| 2266 | and f': "(f has_vector_derivative f') (at x within s)" | |
| 2267 | and f'': "(f has_vector_derivative f'') (at x within s)" | |
| 37730 | 2268 | shows "f' = f''" | 
| 53781 | 2269 | proof - | 
| 37730 | 2270 | have "(\<lambda>x. x *\<^sub>R f') = (\<lambda>x. x *\<^sub>R f'')" | 
| 61245 | 2271 | proof (rule frechet_derivative_unique_within) | 
| 2272 | show "\<forall>i\<in>Basis. \<forall>e>0. \<exists>d. 0 < \<bar>d\<bar> \<and> \<bar>d\<bar> < e \<and> x + d *\<^sub>R i \<in> s" | |
| 2273 | proof clarsimp | |
| 2274 | fix e :: real assume "0 < e" | |
| 2275 | with islimpt_approachable_real[of x s] not_bot | |
| 2276 | obtain x' where "x' \<in> s" "x' \<noteq> x" "\<bar>x' - x\<bar> < e" | |
| 2277 | by (auto simp add: trivial_limit_within) | |
| 2278 | then show "\<exists>d. d \<noteq> 0 \<and> \<bar>d\<bar> < e \<and> x + d \<in> s" | |
| 2279 | by (intro exI[of _ "x' - x"]) auto | |
| 2280 | qed | |
| 2281 | qed (insert f' f'', auto simp: has_vector_derivative_def) | |
| 53781 | 2282 | then show ?thesis | 
| 61649 
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
 paulson <lp15@cam.ac.uk> parents: 
61560diff
changeset | 2283 | unfolding fun_eq_iff by (metis scaleR_one) | 
| 37730 | 2284 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2285 | |
| 61245 | 2286 | lemma vector_derivative_unique_at: | 
| 2287 | "(f has_vector_derivative f') (at x) \<Longrightarrow> (f has_vector_derivative f'') (at x) \<Longrightarrow> f' = f''" | |
| 2288 | by (rule vector_derivative_unique_within) auto | |
| 2289 | ||
| 2290 | lemma differentiableI_vector: "(f has_vector_derivative y) F \<Longrightarrow> f differentiable F" | |
| 2291 | by (auto simp: differentiable_def has_vector_derivative_def) | |
| 2292 | ||
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 2293 | lemma vector_derivative_works: | 
| 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 2294 | "f differentiable net \<longleftrightarrow> (f has_vector_derivative (vector_derivative f net)) net" | 
| 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 2295 | (is "?l = ?r") | 
| 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 2296 | proof | 
| 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 2297 | assume ?l | 
| 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 2298 | obtain f' where f': "(f has_derivative f') net" | 
| 60420 | 2299 | using \<open>?l\<close> unfolding differentiable_def .. | 
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 2300 | then interpret bounded_linear f' | 
| 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 2301 | by auto | 
| 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 2302 | show ?r | 
| 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 2303 | unfolding vector_derivative_def has_vector_derivative_def | 
| 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 2304 | by (rule someI[of _ "f' 1"]) (simp add: scaleR[symmetric] f') | 
| 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 2305 | qed (auto simp: vector_derivative_def has_vector_derivative_def differentiable_def) | 
| 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 2306 | |
| 61245 | 2307 | lemma vector_derivative_within: | 
| 2308 | assumes not_bot: "at x within s \<noteq> bot" and y: "(f has_vector_derivative y) (at x within s)" | |
| 2309 | shows "vector_derivative f (at x within s) = y" | |
| 2310 | using y | |
| 2311 | by (intro vector_derivative_unique_within[OF not_bot vector_derivative_works[THEN iffD1] y]) | |
| 2312 | (auto simp: differentiable_def has_vector_derivative_def) | |
| 2313 | ||
| 61520 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 paulson <lp15@cam.ac.uk> parents: 
61518diff
changeset | 2314 | lemma frechet_derivative_eq_vector_derivative: | 
| 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 paulson <lp15@cam.ac.uk> parents: 
61518diff
changeset | 2315 | assumes "f differentiable (at x)" | 
| 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 paulson <lp15@cam.ac.uk> parents: 
61518diff
changeset | 2316 | shows "(frechet_derivative f (at x)) = (\<lambda>r. r *\<^sub>R vector_derivative f (at x))" | 
| 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 paulson <lp15@cam.ac.uk> parents: 
61518diff
changeset | 2317 | using assms | 
| 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 paulson <lp15@cam.ac.uk> parents: 
61518diff
changeset | 2318 | by (auto simp: differentiable_iff_scaleR vector_derivative_def has_vector_derivative_def | 
| 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 paulson <lp15@cam.ac.uk> parents: 
61518diff
changeset | 2319 | intro: someI frechet_derivative_at [symmetric]) | 
| 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 paulson <lp15@cam.ac.uk> parents: 
61518diff
changeset | 2320 | |
| 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 paulson <lp15@cam.ac.uk> parents: 
61518diff
changeset | 2321 | lemma has_real_derivative: | 
| 61649 
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
 paulson <lp15@cam.ac.uk> parents: 
61560diff
changeset | 2322 | fixes f :: "real \<Rightarrow> real" | 
| 61520 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 paulson <lp15@cam.ac.uk> parents: 
61518diff
changeset | 2323 | assumes "(f has_derivative f') F" | 
| 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 paulson <lp15@cam.ac.uk> parents: 
61518diff
changeset | 2324 | obtains c where "(f has_real_derivative c) F" | 
| 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 paulson <lp15@cam.ac.uk> parents: 
61518diff
changeset | 2325 | proof - | 
| 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 paulson <lp15@cam.ac.uk> parents: 
61518diff
changeset | 2326 | obtain c where "f' = (\<lambda>x. x * c)" | 
| 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 paulson <lp15@cam.ac.uk> parents: 
61518diff
changeset | 2327 | by (metis assms has_derivative_bounded_linear real_bounded_linear) | 
| 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 paulson <lp15@cam.ac.uk> parents: 
61518diff
changeset | 2328 | then show ?thesis | 
| 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 paulson <lp15@cam.ac.uk> parents: 
61518diff
changeset | 2329 | by (metis assms that has_field_derivative_def mult_commute_abs) | 
| 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 paulson <lp15@cam.ac.uk> parents: 
61518diff
changeset | 2330 | qed | 
| 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 paulson <lp15@cam.ac.uk> parents: 
61518diff
changeset | 2331 | |
| 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 paulson <lp15@cam.ac.uk> parents: 
61518diff
changeset | 2332 | lemma has_real_derivative_iff: | 
| 61649 
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
 paulson <lp15@cam.ac.uk> parents: 
61560diff
changeset | 2333 | fixes f :: "real \<Rightarrow> real" | 
| 61520 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 paulson <lp15@cam.ac.uk> parents: 
61518diff
changeset | 2334 | shows "(\<exists>c. (f has_real_derivative c) F) = (\<exists>D. (f has_derivative D) F)" | 
| 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 paulson <lp15@cam.ac.uk> parents: 
61518diff
changeset | 2335 | by (metis has_field_derivative_def has_real_derivative) | 
| 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 paulson <lp15@cam.ac.uk> parents: 
61518diff
changeset | 2336 | |
| 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 paulson <lp15@cam.ac.uk> parents: 
61518diff
changeset | 2337 | definition deriv :: "('a \<Rightarrow> 'a::real_normed_field) \<Rightarrow> 'a \<Rightarrow> 'a" where
 | 
| 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 paulson <lp15@cam.ac.uk> parents: 
61518diff
changeset | 2338 | "deriv f x \<equiv> SOME D. DERIV f x :> D" | 
| 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 paulson <lp15@cam.ac.uk> parents: 
61518diff
changeset | 2339 | |
| 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 paulson <lp15@cam.ac.uk> parents: 
61518diff
changeset | 2340 | lemma DERIV_imp_deriv: "DERIV f x :> f' \<Longrightarrow> deriv f x = f'" | 
| 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 paulson <lp15@cam.ac.uk> parents: 
61518diff
changeset | 2341 | unfolding deriv_def by (metis some_equality DERIV_unique) | 
| 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 paulson <lp15@cam.ac.uk> parents: 
61518diff
changeset | 2342 | |
| 61907 
f0c894ab18c9
Liouville theorem, Fundamental Theorem of Algebra, etc.
 paulson <lp15@cam.ac.uk> parents: 
61880diff
changeset | 2343 | lemma DERIV_deriv_iff_has_field_derivative: | 
| 
f0c894ab18c9
Liouville theorem, Fundamental Theorem of Algebra, etc.
 paulson <lp15@cam.ac.uk> parents: 
61880diff
changeset | 2344 | "DERIV f x :> deriv f x \<longleftrightarrow> (\<exists>f'. (f has_field_derivative f') (at x))" | 
| 
f0c894ab18c9
Liouville theorem, Fundamental Theorem of Algebra, etc.
 paulson <lp15@cam.ac.uk> parents: 
61880diff
changeset | 2345 | by (auto simp: has_field_derivative_def DERIV_imp_deriv) | 
| 
f0c894ab18c9
Liouville theorem, Fundamental Theorem of Algebra, etc.
 paulson <lp15@cam.ac.uk> parents: 
61880diff
changeset | 2346 | |
| 61520 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 paulson <lp15@cam.ac.uk> parents: 
61518diff
changeset | 2347 | lemma DERIV_deriv_iff_real_differentiable: | 
| 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 paulson <lp15@cam.ac.uk> parents: 
61518diff
changeset | 2348 | fixes x :: real | 
| 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 paulson <lp15@cam.ac.uk> parents: 
61518diff
changeset | 2349 | shows "DERIV f x :> deriv f x \<longleftrightarrow> f differentiable at x" | 
| 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 paulson <lp15@cam.ac.uk> parents: 
61518diff
changeset | 2350 | unfolding differentiable_def by (metis DERIV_imp_deriv has_real_derivative_iff) | 
| 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 paulson <lp15@cam.ac.uk> parents: 
61518diff
changeset | 2351 | |
| 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 paulson <lp15@cam.ac.uk> parents: 
61518diff
changeset | 2352 | lemma real_derivative_chain: | 
| 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 paulson <lp15@cam.ac.uk> parents: 
61518diff
changeset | 2353 | fixes x :: real | 
| 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 paulson <lp15@cam.ac.uk> parents: 
61518diff
changeset | 2354 | shows "f differentiable at x \<Longrightarrow> g differentiable at (f x) | 
| 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 paulson <lp15@cam.ac.uk> parents: 
61518diff
changeset | 2355 | \<Longrightarrow> deriv (g o f) x = deriv g (f x) * deriv f x" | 
| 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 paulson <lp15@cam.ac.uk> parents: 
61518diff
changeset | 2356 | by (metis DERIV_deriv_iff_real_differentiable DERIV_chain DERIV_imp_deriv) | 
| 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 paulson <lp15@cam.ac.uk> parents: 
61518diff
changeset | 2357 | lemma field_derivative_eq_vector_derivative: | 
| 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 paulson <lp15@cam.ac.uk> parents: 
61518diff
changeset | 2358 | "(deriv f x) = vector_derivative f (at x)" | 
| 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 paulson <lp15@cam.ac.uk> parents: 
61518diff
changeset | 2359 | by (simp add: mult.commute deriv_def vector_derivative_def has_vector_derivative_def has_field_derivative_def) | 
| 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 paulson <lp15@cam.ac.uk> parents: 
61518diff
changeset | 2360 | |
| 61245 | 2361 | lemma islimpt_closure_open: | 
| 2362 | fixes s :: "'a::perfect_space set" | |
| 2363 | assumes "open s" and t: "t = closure s" "x \<in> t" | |
| 2364 | shows "x islimpt t" | |
| 2365 | proof cases | |
| 61649 
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
 paulson <lp15@cam.ac.uk> parents: 
61560diff
changeset | 2366 | assume "x \<in> s" | 
| 61245 | 2367 |   { fix T assume "x \<in> T" "open T"
 | 
| 2368 | then have "open (s \<inter> T)" | |
| 2369 | using \<open>open s\<close> by auto | |
| 2370 |     then have "s \<inter> T \<noteq> {x}"
 | |
| 2371 | using not_open_singleton[of x] by auto | |
| 2372 | with \<open>x \<in> T\<close> \<open>x \<in> s\<close> have "\<exists>y\<in>t. y \<in> T \<and> y \<noteq> x" | |
| 2373 | using closure_subset[of s] by (auto simp: t) } | |
| 2374 | then show ?thesis | |
| 2375 | by (auto intro!: islimptI) | |
| 2376 | next | |
| 2377 | assume "x \<notin> s" with t show ?thesis | |
| 2378 | unfolding t closure_def by (auto intro: islimpt_subset) | |
| 2379 | qed | |
| 2380 | ||
| 44123 | 2381 | lemma vector_derivative_unique_within_closed_interval: | 
| 61245 | 2382 | assumes ab: "a < b" "x \<in> cbox a b" | 
| 2383 | assumes D: "(f has_vector_derivative f') (at x within cbox a b)" "(f has_vector_derivative f'') (at x within cbox a b)" | |
| 44123 | 2384 | shows "f' = f''" | 
| 61245 | 2385 | using ab | 
| 2386 | by (intro vector_derivative_unique_within[OF _ D]) | |
| 2387 |      (auto simp: trivial_limit_within intro!: islimpt_closure_open[where s="{a <..< b}"])
 | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2388 | |
| 37730 | 2389 | lemma vector_derivative_at: | 
| 53781 | 2390 | "(f has_vector_derivative f') (at x) \<Longrightarrow> vector_derivative f (at x) = f'" | 
| 61245 | 2391 | by (intro vector_derivative_within at_neq_bot) | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2392 | |
| 61104 
3c2d4636cebc
new lemmas about vector_derivative, complex numbers, paths, etc.
 paulson parents: 
61076diff
changeset | 2393 | lemma has_vector_derivative_id_at [simp]: "vector_derivative (\<lambda>x. x) (at a) = 1" | 
| 
3c2d4636cebc
new lemmas about vector_derivative, complex numbers, paths, etc.
 paulson parents: 
61076diff
changeset | 2394 | by (simp add: vector_derivative_at) | 
| 
3c2d4636cebc
new lemmas about vector_derivative, complex numbers, paths, etc.
 paulson parents: 
61076diff
changeset | 2395 | |
| 
3c2d4636cebc
new lemmas about vector_derivative, complex numbers, paths, etc.
 paulson parents: 
61076diff
changeset | 2396 | lemma vector_derivative_minus_at [simp]: | 
| 
3c2d4636cebc
new lemmas about vector_derivative, complex numbers, paths, etc.
 paulson parents: 
61076diff
changeset | 2397 | "f differentiable at a | 
| 
3c2d4636cebc
new lemmas about vector_derivative, complex numbers, paths, etc.
 paulson parents: 
61076diff
changeset | 2398 | \<Longrightarrow> vector_derivative (\<lambda>x. - f x) (at a) = - vector_derivative f (at a)" | 
| 
3c2d4636cebc
new lemmas about vector_derivative, complex numbers, paths, etc.
 paulson parents: 
61076diff
changeset | 2399 | by (simp add: vector_derivative_at has_vector_derivative_minus vector_derivative_works [symmetric]) | 
| 
3c2d4636cebc
new lemmas about vector_derivative, complex numbers, paths, etc.
 paulson parents: 
61076diff
changeset | 2400 | |
| 
3c2d4636cebc
new lemmas about vector_derivative, complex numbers, paths, etc.
 paulson parents: 
61076diff
changeset | 2401 | lemma vector_derivative_add_at [simp]: | 
| 
3c2d4636cebc
new lemmas about vector_derivative, complex numbers, paths, etc.
 paulson parents: 
61076diff
changeset | 2402 | "\<lbrakk>f differentiable at a; g differentiable at a\<rbrakk> | 
| 
3c2d4636cebc
new lemmas about vector_derivative, complex numbers, paths, etc.
 paulson parents: 
61076diff
changeset | 2403 | \<Longrightarrow> vector_derivative (\<lambda>x. f x + g x) (at a) = vector_derivative f (at a) + vector_derivative g (at a)" | 
| 
3c2d4636cebc
new lemmas about vector_derivative, complex numbers, paths, etc.
 paulson parents: 
61076diff
changeset | 2404 | by (simp add: vector_derivative_at has_vector_derivative_add vector_derivative_works [symmetric]) | 
| 
3c2d4636cebc
new lemmas about vector_derivative, complex numbers, paths, etc.
 paulson parents: 
61076diff
changeset | 2405 | |
| 
3c2d4636cebc
new lemmas about vector_derivative, complex numbers, paths, etc.
 paulson parents: 
61076diff
changeset | 2406 | lemma vector_derivative_diff_at [simp]: | 
| 
3c2d4636cebc
new lemmas about vector_derivative, complex numbers, paths, etc.
 paulson parents: 
61076diff
changeset | 2407 | "\<lbrakk>f differentiable at a; g differentiable at a\<rbrakk> | 
| 
3c2d4636cebc
new lemmas about vector_derivative, complex numbers, paths, etc.
 paulson parents: 
61076diff
changeset | 2408 | \<Longrightarrow> vector_derivative (\<lambda>x. f x - g x) (at a) = vector_derivative f (at a) - vector_derivative g (at a)" | 
| 
3c2d4636cebc
new lemmas about vector_derivative, complex numbers, paths, etc.
 paulson parents: 
61076diff
changeset | 2409 | by (simp add: vector_derivative_at has_vector_derivative_diff vector_derivative_works [symmetric]) | 
| 
3c2d4636cebc
new lemmas about vector_derivative, complex numbers, paths, etc.
 paulson parents: 
61076diff
changeset | 2410 | |
| 61204 | 2411 | lemma vector_derivative_mult_at [simp]: | 
| 2412 | fixes f g :: "real \<Rightarrow> 'a :: real_normed_algebra" | |
| 2413 | shows "\<lbrakk>f differentiable at a; g differentiable at a\<rbrakk> | |
| 2414 | \<Longrightarrow> vector_derivative (\<lambda>x. f x * g x) (at a) = f a * vector_derivative g (at a) + vector_derivative f (at a) * g a" | |
| 2415 | by (simp add: vector_derivative_at has_vector_derivative_mult vector_derivative_works [symmetric]) | |
| 2416 | ||
| 2417 | lemma vector_derivative_scaleR_at [simp]: | |
| 2418 | "\<lbrakk>f differentiable at a; g differentiable at a\<rbrakk> | |
| 2419 | \<Longrightarrow> vector_derivative (\<lambda>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" | |
| 2420 | apply (rule vector_derivative_at) | |
| 2421 | apply (rule has_vector_derivative_scaleR) | |
| 2422 | apply (auto simp: vector_derivative_works has_vector_derivative_def has_field_derivative_def mult_commute_abs) | |
| 2423 | done | |
| 2424 | ||
| 44123 | 2425 | lemma vector_derivative_within_closed_interval: | 
| 61245 | 2426 | assumes ab: "a < b" "x \<in> cbox a b" | 
| 2427 | assumes f: "(f has_vector_derivative f') (at x within cbox a b)" | |
| 56188 | 2428 | shows "vector_derivative f (at x within cbox a b) = f'" | 
| 61245 | 2429 | by (intro vector_derivative_unique_within_closed_interval[OF ab _ f] | 
| 2430 | vector_derivative_works[THEN iffD1] differentiableI_vector) | |
| 2431 | fact | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2432 | |
| 53781 | 2433 | lemma has_vector_derivative_within_subset: | 
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 2434 | "(f has_vector_derivative f') (at x within s) \<Longrightarrow> t \<subseteq> s \<Longrightarrow> (f has_vector_derivative f') (at x within t)" | 
| 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 2435 | by (auto simp: has_vector_derivative_def intro: has_derivative_within_subset) | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2436 | |
| 44123 | 2437 | lemma has_vector_derivative_at_within: | 
| 2438 | "(f has_vector_derivative f') (at x) \<Longrightarrow> (f has_vector_derivative f') (at x within s)" | |
| 2439 | unfolding has_vector_derivative_def | |
| 45031 | 2440 | by (rule has_derivative_at_within) | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2441 | |
| 61880 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61824diff
changeset | 2442 | lemma has_vector_derivative_weaken: | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61824diff
changeset | 2443 | fixes x D and f g s t | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61824diff
changeset | 2444 | assumes f: "(f has_vector_derivative D) (at x within t)" | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61824diff
changeset | 2445 | and "x \<in> s" "s \<subseteq> t" | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61824diff
changeset | 2446 | and "\<And>x. x \<in> s \<Longrightarrow> f x = g x" | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61824diff
changeset | 2447 | shows "(g has_vector_derivative D) (at x within s)" | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61824diff
changeset | 2448 | proof - | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61824diff
changeset | 2449 | have "(f has_vector_derivative D) (at x within s) \<longleftrightarrow> (g has_vector_derivative D) (at x within s)" | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61824diff
changeset | 2450 | unfolding has_vector_derivative_def has_derivative_iff_norm | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61824diff
changeset | 2451 | using assms by (intro conj_cong Lim_cong_within refl) auto | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61824diff
changeset | 2452 | then show ?thesis | 
| 61975 | 2453 | using has_vector_derivative_within_subset[OF f \<open>s \<subseteq> t\<close>] by simp | 
| 61880 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61824diff
changeset | 2454 | qed | 
| 
ff4d33058566
moved some theorems from the CLT proof; reordered some theorems / notation
 hoelzl parents: 
61824diff
changeset | 2455 | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2456 | lemma has_vector_derivative_transform_within: | 
| 62087 
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
 paulson parents: 
61975diff
changeset | 2457 | assumes "(f has_vector_derivative f') (at x within s)" | 
| 
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
 paulson parents: 
61975diff
changeset | 2458 | and "0 < d" | 
| 53781 | 2459 | and "x \<in> s" | 
| 62087 
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
 paulson parents: 
61975diff
changeset | 2460 | and "\<And>x'. \<lbrakk>x'\<in>s; dist x' x < d\<rbrakk> \<Longrightarrow> f x' = g x'" | 
| 
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
 paulson parents: 
61975diff
changeset | 2461 | shows "(g has_vector_derivative f') (at x within s)" | 
| 53781 | 2462 | using assms | 
| 2463 | unfolding has_vector_derivative_def | |
| 44123 | 2464 | by (rule has_derivative_transform_within) | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2465 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2466 | lemma has_vector_derivative_transform_within_open: | 
| 62087 
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
 paulson parents: 
61975diff
changeset | 2467 | assumes "(f has_vector_derivative f') (at x)" | 
| 
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
 paulson parents: 
61975diff
changeset | 2468 | and "open s" | 
| 53781 | 2469 | and "x \<in> s" | 
| 62087 
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
 paulson parents: 
61975diff
changeset | 2470 | and "\<And>y. y\<in>s \<Longrightarrow> f y = g y" | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2471 | shows "(g has_vector_derivative f') (at x)" | 
| 53781 | 2472 | using assms | 
| 2473 | unfolding has_vector_derivative_def | |
| 44123 | 2474 | by (rule has_derivative_transform_within_open) | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2475 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2476 | lemma vector_diff_chain_at: | 
| 44123 | 2477 | assumes "(f has_vector_derivative f') (at x)" | 
| 53781 | 2478 | and "(g has_vector_derivative g') (at (f x))" | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2479 | shows "((g \<circ> f) has_vector_derivative (f' *\<^sub>R g')) (at x)" | 
| 53781 | 2480 | using assms(2) | 
| 2481 | unfolding has_vector_derivative_def | |
| 2482 | apply - | |
| 2483 | apply (drule diff_chain_at[OF assms(1)[unfolded has_vector_derivative_def]]) | |
| 2484 | apply (simp only: o_def real_scaleR_def scaleR_scaleR) | |
| 2485 | done | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2486 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2487 | lemma vector_diff_chain_within: | 
| 44123 | 2488 | assumes "(f has_vector_derivative f') (at x within s)" | 
| 53781 | 2489 | and "(g has_vector_derivative g') (at (f x) within f ` s)" | 
| 2490 | shows "((g \<circ> f) has_vector_derivative (f' *\<^sub>R g')) (at x within s)" | |
| 2491 | using assms(2) | |
| 2492 | unfolding has_vector_derivative_def | |
| 2493 | apply - | |
| 2494 | apply (drule diff_chain_within[OF assms(1)[unfolded has_vector_derivative_def]]) | |
| 2495 | apply (simp only: o_def real_scaleR_def scaleR_scaleR) | |
| 2496 | done | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2497 | |
| 60762 | 2498 | lemma vector_derivative_const_at [simp]: "vector_derivative (\<lambda>x. c) (at a) = 0" | 
| 2499 | by (simp add: vector_derivative_at) | |
| 2500 | ||
| 60800 
7d04351c795a
New material for Cauchy's integral theorem
 paulson <lp15@cam.ac.uk> parents: 
60762diff
changeset | 2501 | lemma vector_derivative_at_within_ivl: | 
| 
7d04351c795a
New material for Cauchy's integral theorem
 paulson <lp15@cam.ac.uk> parents: 
60762diff
changeset | 2502 | "(f has_vector_derivative f') (at x) \<Longrightarrow> | 
| 
7d04351c795a
New material for Cauchy's integral theorem
 paulson <lp15@cam.ac.uk> parents: 
60762diff
changeset | 2503 |     a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> a<b \<Longrightarrow> vector_derivative f (at x within {a..b}) = f'"
 | 
| 
7d04351c795a
New material for Cauchy's integral theorem
 paulson <lp15@cam.ac.uk> parents: 
60762diff
changeset | 2504 | using has_vector_derivative_at_within vector_derivative_within_closed_interval by fastforce | 
| 
7d04351c795a
New material for Cauchy's integral theorem
 paulson <lp15@cam.ac.uk> parents: 
60762diff
changeset | 2505 | |
| 61204 | 2506 | lemma vector_derivative_chain_at: | 
| 2507 | assumes "f differentiable at x" "(g differentiable at (f x))" | |
| 2508 | shows "vector_derivative (g \<circ> f) (at x) = | |
| 2509 | vector_derivative f (at x) *\<^sub>R vector_derivative g (at (f x))" | |
| 2510 | by (metis vector_diff_chain_at vector_derivative_at vector_derivative_works assms) | |
| 2511 | ||
| 62408 
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
 paulson <lp15@cam.ac.uk> parents: 
62393diff
changeset | 2512 | lemma field_vector_diff_chain_at: (*thanks to Wenda Li*) | 
| 
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
 paulson <lp15@cam.ac.uk> parents: 
62393diff
changeset | 2513 | assumes Df: "(f has_vector_derivative f') (at x)" | 
| 
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
 paulson <lp15@cam.ac.uk> parents: 
62393diff
changeset | 2514 | and Dg: "(g has_field_derivative g') (at (f x))" | 
| 
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
 paulson <lp15@cam.ac.uk> parents: 
62393diff
changeset | 2515 | shows "((g \<circ> f) has_vector_derivative (f' * g')) (at x)" | 
| 
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
 paulson <lp15@cam.ac.uk> parents: 
62393diff
changeset | 2516 | using diff_chain_at[OF Df[unfolded has_vector_derivative_def] | 
| 
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
 paulson <lp15@cam.ac.uk> parents: 
62393diff
changeset | 2517 | Dg [unfolded has_field_derivative_def]] | 
| 
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
 paulson <lp15@cam.ac.uk> parents: 
62393diff
changeset | 2518 | by (auto simp: o_def mult.commute has_vector_derivative_def) | 
| 
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
 paulson <lp15@cam.ac.uk> parents: 
62393diff
changeset | 2519 | |
| 
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
 paulson <lp15@cam.ac.uk> parents: 
62393diff
changeset | 2520 | lemma vector_derivative_chain_at_general: (*thanks to Wenda Li*) | 
| 
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
 paulson <lp15@cam.ac.uk> parents: 
62393diff
changeset | 2521 | assumes "f differentiable at x" "(\<exists>g'. (g has_field_derivative g') (at (f x)))" | 
| 
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
 paulson <lp15@cam.ac.uk> parents: 
62393diff
changeset | 2522 | shows "vector_derivative (g \<circ> f) (at x) = | 
| 
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
 paulson <lp15@cam.ac.uk> parents: 
62393diff
changeset | 2523 | vector_derivative f (at x) * deriv g (f x)" | 
| 
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
 paulson <lp15@cam.ac.uk> parents: 
62393diff
changeset | 2524 | apply (rule vector_derivative_at [OF field_vector_diff_chain_at]) | 
| 
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
 paulson <lp15@cam.ac.uk> parents: 
62393diff
changeset | 2525 | using assms | 
| 
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
 paulson <lp15@cam.ac.uk> parents: 
62393diff
changeset | 2526 | by (auto simp: vector_derivative_works DERIV_deriv_iff_has_field_derivative) | 
| 
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
 paulson <lp15@cam.ac.uk> parents: 
62393diff
changeset | 2527 | |
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2528 | |
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2529 | subsection \<open>Relation between convexity and derivative\<close> | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2530 | |
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2531 | (* TODO: Generalise to real vector spaces? *) | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2532 | lemma convex_on_imp_above_tangent: | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2533 | assumes convex: "convex_on A f" and connected: "connected A" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2534 | assumes c: "c \<in> interior A" and x : "x \<in> A" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2535 | assumes deriv: "(f has_field_derivative f') (at c within A)" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2536 | shows "f x - f c \<ge> f' * (x - c)" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2537 | proof (cases x c rule: linorder_cases) | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2538 | assume xc: "x > c" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2539 |   let ?A' = "interior A \<inter> {c<..}"
 | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2540 |   from c have "c \<in> interior A \<inter> closure {c<..}" by auto
 | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2541 |   also have "\<dots> \<subseteq> closure (interior A \<inter> {c<..})" by (intro open_inter_closure_subset) auto
 | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2542 | finally have "at c within ?A' \<noteq> bot" by (subst at_within_eq_bot_iff) auto | 
| 61973 | 2543 | moreover from deriv have "((\<lambda>y. (f y - f c) / (y - c)) \<longlongrightarrow> f') (at c within ?A')" | 
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2544 | unfolding DERIV_within_iff using interior_subset[of A] by (blast intro: tendsto_mono at_le) | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2545 | moreover from eventually_at_right_real[OF xc] | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2546 | have "eventually (\<lambda>y. (f y - f c) / (y - c) \<le> (f x - f c) / (x - c)) (at_right c)" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2547 | proof eventually_elim | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2548 |     fix y assume y: "y \<in> {c<..<x}"
 | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2549 | with convex connected x c have "f y \<le> (f x - f c) / (x - c) * (y - c) + f c" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2550 | using interior_subset[of A] | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2551 | by (intro convex_onD_Icc' convex_on_subset[OF convex] connected_contains_Icc) auto | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2552 | hence "f y - f c \<le> (f x - f c) / (x - c) * (y - c)" by simp | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2553 | thus "(f y - f c) / (y - c) \<le> (f x - f c) / (x - c)" using y xc by (simp add: divide_simps) | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2554 | qed | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2555 | hence "eventually (\<lambda>y. (f y - f c) / (y - c) \<le> (f x - f c) / (x - c)) (at c within ?A')" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2556 | by (blast intro: filter_leD at_le) | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2557 | ultimately have "f' \<le> (f x - f c) / (x - c)" by (rule tendsto_ge_const) | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2558 | thus ?thesis using xc by (simp add: field_simps) | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2559 | next | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2560 | assume xc: "x < c" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2561 |   let ?A' = "interior A \<inter> {..<c}"
 | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2562 |   from c have "c \<in> interior A \<inter> closure {..<c}" by auto
 | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2563 |   also have "\<dots> \<subseteq> closure (interior A \<inter> {..<c})" by (intro open_inter_closure_subset) auto
 | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2564 | finally have "at c within ?A' \<noteq> bot" by (subst at_within_eq_bot_iff) auto | 
| 61973 | 2565 | moreover from deriv have "((\<lambda>y. (f y - f c) / (y - c)) \<longlongrightarrow> f') (at c within ?A')" | 
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2566 | unfolding DERIV_within_iff using interior_subset[of A] by (blast intro: tendsto_mono at_le) | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2567 | moreover from eventually_at_left_real[OF xc] | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2568 | have "eventually (\<lambda>y. (f y - f c) / (y - c) \<ge> (f x - f c) / (x - c)) (at_left c)" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2569 | proof eventually_elim | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2570 |     fix y assume y: "y \<in> {x<..<c}"
 | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2571 | with convex connected x c have "f y \<le> (f x - f c) / (c - x) * (c - y) + f c" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2572 | using interior_subset[of A] | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2573 | by (intro convex_onD_Icc'' convex_on_subset[OF convex] connected_contains_Icc) auto | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2574 | hence "f y - f c \<le> (f x - f c) * ((c - y) / (c - x))" by simp | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2575 | also have "(c - y) / (c - x) = (y - c) / (x - c)" using y xc by (simp add: field_simps) | 
| 61649 
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
 paulson <lp15@cam.ac.uk> parents: 
61560diff
changeset | 2576 | finally show "(f y - f c) / (y - c) \<ge> (f x - f c) / (x - c)" using y xc | 
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2577 | by (simp add: divide_simps) | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2578 | qed | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2579 | hence "eventually (\<lambda>y. (f y - f c) / (y - c) \<ge> (f x - f c) / (x - c)) (at c within ?A')" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2580 | by (blast intro: filter_leD at_le) | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2581 | ultimately have "f' \<ge> (f x - f c) / (x - c)" by (rule tendsto_le_const) | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2582 | thus ?thesis using xc by (simp add: field_simps) | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2583 | qed simp_all | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61524diff
changeset | 2584 | |
| 62207 | 2585 | |
| 2586 | subsection \<open>Partial derivatives\<close> | |
| 2587 | ||
| 2588 | lemma eventually_at_Pair_within_TimesI1: | |
| 2589 | fixes x::"'a::metric_space" | |
| 2590 | assumes "\<forall>\<^sub>F x' in at x within X. P x'" | |
| 2591 | assumes "P x" | |
| 2592 | shows "\<forall>\<^sub>F (x', y') in at (x, y) within X \<times> Y. P x'" | |
| 2593 | proof - | |
| 2594 | from assms[unfolded eventually_at_topological] | |
| 2595 | obtain S where S: "open S" "x \<in> S" "\<And>x'. x' \<in> X \<Longrightarrow> x' \<in> S \<Longrightarrow> P x'" | |
| 2596 | by metis | |
| 2597 | show "\<forall>\<^sub>F (x', y') in at (x, y) within X \<times> Y. P x'" | |
| 2598 | unfolding eventually_at_topological | |
| 2599 | by (auto intro!: exI[where x="S \<times> UNIV"] S open_Times) | |
| 2600 | qed | |
| 2601 | ||
| 2602 | lemma eventually_at_Pair_within_TimesI2: | |
| 2603 | fixes x::"'a::metric_space" | |
| 2604 | assumes "\<forall>\<^sub>F y' in at y within Y. P y'" | |
| 2605 | assumes "P y" | |
| 2606 | shows "\<forall>\<^sub>F (x', y') in at (x, y) within X \<times> Y. P y'" | |
| 2607 | proof - | |
| 2608 | from assms[unfolded eventually_at_topological] | |
| 2609 | obtain S where S: "open S" "y \<in> S" "\<And>y'. y' \<in> Y \<Longrightarrow> y' \<in> S \<Longrightarrow> P y'" | |
| 2610 | by metis | |
| 2611 | show "\<forall>\<^sub>F (x', y') in at (x, y) within X \<times> Y. P y'" | |
| 2612 | unfolding eventually_at_topological | |
| 2613 | by (auto intro!: exI[where x="UNIV \<times> S"] S open_Times) | |
| 2614 | qed | |
| 2615 | ||
| 2616 | lemma has_derivative_partialsI: | |
| 2617 | assumes fx: "\<And>x y. x \<in> X \<Longrightarrow> y \<in> Y \<Longrightarrow> ((\<lambda>x. f x y) has_derivative blinfun_apply (fx x y)) (at x within X)" | |
| 2618 | assumes fy: "\<And>x y. x \<in> X \<Longrightarrow> y \<in> Y \<Longrightarrow> ((\<lambda>y. f x y) has_derivative blinfun_apply (fy x y)) (at y within Y)" | |
| 2619 | assumes fx_cont: "continuous_on (X \<times> Y) (\<lambda>(x, y). fx x y)" | |
| 2620 | assumes fy_cont: "continuous_on (X \<times> Y) (\<lambda>(x, y). fy x y)" | |
| 2621 | assumes "x \<in> X" "y \<in> Y" | |
| 2622 | assumes "convex X" "convex Y" | |
| 2623 | shows "((\<lambda>(x, y). f x y) has_derivative (\<lambda>(tx, ty). fx x y tx + fy x y ty)) (at (x, y) within X \<times> Y)" | |
| 2624 | proof (safe intro!: has_derivativeI tendstoI, goal_cases) | |
| 2625 | case (2 e') | |
| 2626 | def e\<equiv>"e' / 9" | |
| 2627 | have "e > 0" using \<open>e' > 0\<close> by (simp add: e_def) | |
| 2628 | ||
| 2629 | have "(x, y) \<in> X \<times> Y" using assms by auto | |
| 2630 | from fy_cont[unfolded continuous_on_eq_continuous_within, rule_format, OF this, | |
| 2631 | unfolded continuous_within, THEN tendstoD, OF \<open>e > 0\<close>] | |
| 2632 | have "\<forall>\<^sub>F (x', y') in at (x, y) within X \<times> Y. dist (fy x' y') (fy x y) < e" | |
| 2633 | by (auto simp: split_beta') | |
| 2634 | from this[unfolded eventually_at] obtain d' where | |
| 2635 | "d' > 0" | |
| 2636 | "\<And>x' y'. x' \<in> X \<Longrightarrow> y' \<in> Y \<Longrightarrow> (x', y') \<noteq> (x, y) \<Longrightarrow> dist (x', y') (x, y) < d' \<Longrightarrow> | |
| 2637 | dist (fy x' y') (fy x y) < e" | |
| 2638 | by auto | |
| 2639 | then | |
| 2640 | have d': "x' \<in> X \<Longrightarrow> y' \<in> Y \<Longrightarrow> dist (x', y') (x, y) < d' \<Longrightarrow> dist (fy x' y') (fy x y) < e" | |
| 2641 | for x' y' | |
| 2642 | using \<open>0 < e\<close> | |
| 2643 | by (cases "(x', y') = (x, y)") auto | |
| 2644 | def d \<equiv> "d' / sqrt 2" | |
| 2645 | have "d > 0" using \<open>0 < d'\<close> by (simp add: d_def) | |
| 2646 | have d: "x' \<in> X \<Longrightarrow> y' \<in> Y \<Longrightarrow> dist x' x < d \<Longrightarrow> dist y' y < d \<Longrightarrow> dist (fy x' y') (fy x y) < e" | |
| 2647 | for x' y' | |
| 2648 | by (auto simp: dist_prod_def d_def intro!: d' real_sqrt_sum_squares_less) | |
| 2649 | ||
| 2650 | let ?S = "ball y d \<inter> Y" | |
| 2651 | have "convex ?S" | |
| 2652 | by (auto intro!: convex_Int \<open>convex Y\<close>) | |
| 2653 |   {
 | |
| 2654 | fix x'::'a and y'::'b | |
| 2655 | assume x': "x' \<in> X" and y': "y' \<in> Y" | |
| 2656 | assume dx': "dist x' x < d" and dy': "dist y' y < d" | |
| 2657 | have "norm (fy x' y' - fy x' y) \<le> dist (fy x' y') (fy x y) + dist (fy x' y) (fy x y)" | |
| 2658 | by norm | |
| 2659 | also have "dist (fy x' y') (fy x y) < e" | |
| 2660 | by (rule d; fact) | |
| 2661 | also have "dist (fy x' y) (fy x y) < e" | |
| 2662 | by (auto intro!: d simp: dist_prod_def x' \<open>d > 0\<close> \<open>y \<in> Y\<close> dx') | |
| 2663 | finally | |
| 2664 | have "norm (fy x' y' - fy x' y) < e + e" | |
| 2665 | by arith | |
| 2666 | then have "onorm (blinfun_apply (fy x' y') - blinfun_apply (fy x' y)) < e + e" | |
| 2667 | by (auto simp: norm_blinfun.rep_eq blinfun.diff_left[abs_def] fun_diff_def) | |
| 2668 | } note onorm = this | |
| 2669 | ||
| 2670 | have ev_mem: "\<forall>\<^sub>F (x', y') in at (x, y) within X \<times> Y. (x', y') \<in> X \<times> Y" | |
| 2671 | using \<open>x \<in> X\<close> \<open>y \<in> Y\<close> | |
| 2672 | by (auto simp: eventually_at intro!: zero_less_one) | |
| 2673 | moreover | |
| 2674 | have ev_dist: "\<forall>\<^sub>F xy in at (x, y) within X \<times> Y. dist xy (x, y) < d" if "d > 0" for d | |
| 2675 | using eventually_at_ball[OF that] | |
| 2676 | by (rule eventually_elim2) (auto simp: dist_commute intro!: eventually_True) | |
| 2677 | note ev_dist[OF \<open>0 < d\<close>] | |
| 2678 | ultimately | |
| 2679 | have "\<forall>\<^sub>F (x', y') in at (x, y) within X \<times> Y. | |
| 2680 | norm (f x' y' - f x' y - (fy x' y) (y' - y)) \<le> norm (y' - y) * (e + e)" | |
| 2681 | proof (eventually_elim, safe) | |
| 2682 | fix x' y' | |
| 2683 | assume "x' \<in> X" and y': "y' \<in> Y" | |
| 2684 | assume dist: "dist (x', y') (x, y) < d" | |
| 2685 | then have dx: "dist x' x < d" and dy: "dist y' y < d" | |
| 2686 | unfolding dist_prod_def fst_conv snd_conv atomize_conj | |
| 2687 | by (metis le_less_trans real_sqrt_sum_squares_ge1 real_sqrt_sum_squares_ge2) | |
| 2688 |     {
 | |
| 2689 | fix t::real | |
| 2690 |       assume "t \<in> {0 .. 1}"
 | |
| 2691 | then have "y + t *\<^sub>R (y' - y) \<in> closed_segment y y'" | |
| 2692 | by (auto simp: closed_segment_def algebra_simps intro!: exI[where x=t]) | |
| 2693 | also | |
| 2694 | have "\<dots> \<subseteq> ball y d \<inter> Y" | |
| 2695 | using \<open>y \<in> Y\<close> \<open>0 < d\<close> dy y' | |
| 2696 | by (intro \<open>convex ?S\<close>[unfolded convex_contains_segment, rule_format, of y y']) | |
| 2697 | (auto simp: dist_commute) | |
| 2698 | finally have "y + t *\<^sub>R (y' - y) \<in> ?S" . | |
| 2699 | } note seg = this | |
| 2700 | ||
| 2701 | have "\<forall>x\<in>ball y d \<inter> Y. onorm (blinfun_apply (fy x' x) - blinfun_apply (fy x' y)) \<le> e + e" | |
| 2702 | by (safe intro!: onorm less_imp_le \<open>x' \<in> X\<close> dx) (auto simp: dist_commute \<open>0 < d\<close> \<open>y \<in> Y\<close>) | |
| 2703 | with seg has_derivative_within_subset[OF assms(2)[OF \<open>x' \<in> X\<close>]] | |
| 2704 | show "norm (f x' y' - f x' y - (fy x' y) (y' - y)) \<le> norm (y' - y) * (e + e)" | |
| 2705 | by (rule differentiable_bound_linearization[where S="?S"]) | |
| 2706 | (auto intro!: \<open>0 < d\<close> \<open>y \<in> Y\<close>) | |
| 2707 | qed | |
| 2708 | moreover | |
| 2709 | let ?le = "\<lambda>x'. norm (f x' y - f x y - (fx x y) (x' - x)) \<le> norm (x' - x) * e" | |
| 2710 | from fx[OF \<open>x \<in> X\<close> \<open>y \<in> Y\<close>, unfolded has_derivative_within, THEN conjunct2, THEN tendstoD, OF \<open>0 < e\<close>] | |
| 2711 | have "\<forall>\<^sub>F x' in at x within X. ?le x'" | |
| 2712 | by eventually_elim | |
| 62390 | 2713 | (auto simp: dist_norm divide_simps blinfun.bilinear_simps field_simps split: if_split_asm) | 
| 62207 | 2714 | then have "\<forall>\<^sub>F (x', y') in at (x, y) within X \<times> Y. ?le x'" | 
| 2715 | by (rule eventually_at_Pair_within_TimesI1) | |
| 2716 | (simp add: blinfun.bilinear_simps) | |
| 2717 | moreover have "\<forall>\<^sub>F (x', y') in at (x, y) within X \<times> Y. norm ((x', y') - (x, y)) \<noteq> 0" | |
| 2718 | unfolding norm_eq_zero right_minus_eq | |
| 2719 | by (auto simp: eventually_at intro!: zero_less_one) | |
| 2720 | moreover | |
| 2721 | from fy_cont[unfolded continuous_on_eq_continuous_within, rule_format, OF SigmaI[OF \<open>x \<in> X\<close> \<open>y \<in> Y\<close>], | |
| 2722 | unfolded continuous_within, THEN tendstoD, OF \<open>0 < e\<close>] | |
| 2723 | have "\<forall>\<^sub>F x' in at x within X. norm (fy x' y - fy x y) < e" | |
| 2724 | unfolding eventually_at | |
| 2725 | using \<open>y \<in> Y\<close> | |
| 2726 | by (auto simp: dist_prod_def dist_norm) | |
| 2727 | then have "\<forall>\<^sub>F (x', y') in at (x, y) within X \<times> Y. norm (fy x' y - fy x y) < e" | |
| 2728 | by (rule eventually_at_Pair_within_TimesI1) | |
| 2729 | (simp add: blinfun.bilinear_simps \<open>0 < e\<close>) | |
| 2730 | ultimately | |
| 2731 | have "\<forall>\<^sub>F (x', y') in at (x, y) within X \<times> Y. | |
| 2732 | norm ((f x' y' - f x y - (fx x y (x' - x) + fy x y (y' - y))) /\<^sub>R | |
| 2733 | norm ((x', y') - (x, y))) | |
| 2734 | < e'" | |
| 2735 | apply eventually_elim | |
| 2736 | proof safe | |
| 2737 | fix x' y' | |
| 2738 | have "norm (f x' y' - f x y - (fx x y (x' - x) + fy x y (y' - y))) \<le> | |
| 2739 | norm (f x' y' - f x' y - fy x' y (y' - y)) + | |
| 2740 | norm (fy x y (y' - y) - fy x' y (y' - y)) + | |
| 2741 | norm (f x' y - f x y - fx x y (x' - x))" | |
| 2742 | by norm | |
| 2743 | also | |
| 2744 | assume nz: "norm ((x', y') - (x, y)) \<noteq> 0" | |
| 2745 | and nfy: "norm (fy x' y - fy x y) < e" | |
| 2746 | assume "norm (f x' y' - f x' y - blinfun_apply (fy x' y) (y' - y)) \<le> norm (y' - y) * (e + e)" | |
| 2747 | also assume "norm (f x' y - f x y - blinfun_apply (fx x y) (x' - x)) \<le> norm (x' - x) * e" | |
| 2748 | also | |
| 2749 | have "norm ((fy x y) (y' - y) - (fy x' y) (y' - y)) \<le> norm ((fy x y) - (fy x' y)) * norm (y' - y)" | |
| 2750 | by (auto simp: blinfun.bilinear_simps[symmetric] intro!: norm_blinfun) | |
| 2751 | also have "\<dots> \<le> (e + e) * norm (y' - y)" | |
| 2752 | using \<open>e > 0\<close> nfy | |
| 2753 | by (auto simp: norm_minus_commute intro!: mult_right_mono) | |
| 2754 | also have "norm (x' - x) * e \<le> norm (x' - x) * (e + e)" | |
| 2755 | using \<open>0 < e\<close> by simp | |
| 2756 | also have "norm (y' - y) * (e + e) + (e + e) * norm (y' - y) + norm (x' - x) * (e + e) \<le> | |
| 2757 | (norm (y' - y) + norm (x' - x)) * (4 * e)" | |
| 2758 | using \<open>e > 0\<close> | |
| 2759 | by (simp add: algebra_simps) | |
| 2760 | also have "\<dots> \<le> 2 * norm ((x', y') - (x, y)) * (4 * e)" | |
| 2761 | using \<open>0 < e\<close> real_sqrt_sum_squares_ge1[of "norm (x' - x)" "norm (y' - y)"] | |
| 2762 | real_sqrt_sum_squares_ge2[of "norm (y' - y)" "norm (x' - x)"] | |
| 2763 | by (auto intro!: mult_right_mono simp: norm_prod_def | |
| 2764 | simp del: real_sqrt_sum_squares_ge1 real_sqrt_sum_squares_ge2) | |
| 2765 | also have "\<dots> \<le> norm ((x', y') - (x, y)) * (8 * e)" | |
| 2766 | by simp | |
| 2767 | also have "\<dots> < norm ((x', y') - (x, y)) * e'" | |
| 2768 | using \<open>0 < e'\<close> nz | |
| 2769 | by (auto simp: e_def) | |
| 2770 | finally show "norm ((f x' y' - f x y - (fx x y (x' - x) + fy x y (y' - y))) /\<^sub>R norm ((x', y') - (x, y))) < e'" | |
| 2771 | by (auto simp: divide_simps dist_norm mult.commute) | |
| 2772 | qed | |
| 2773 | then show ?case | |
| 2774 | by eventually_elim (auto simp: dist_norm field_simps) | |
| 2775 | qed (auto intro!: bounded_linear_intros simp: split_beta') | |
| 2776 | ||
| 2777 | ||
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2778 | end |