| author | hoelzl | 
| Tue, 05 Nov 2013 09:44:59 +0100 | |
| changeset 54260 | 6a967667fd45 | 
| parent 54230 | b1d955791529 | 
| child 54775 | 2d3df8633dad | 
| 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 | |
| 53781 | 6 | header {* Multivariate calculus in Euclidean space *}
 | 
| 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 | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36844diff
changeset | 9 | imports Brouwer_Fixpoint Operator_Norm | 
| 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 | |
| 53781 | 12 | lemma bounded_linear_imp_linear: (* TODO: move elsewhere *) | 
| 13 | assumes "bounded_linear f" | |
| 14 | shows "linear f" | |
| 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 | 15 | proof - | 
| 53781 | 16 | interpret f: bounded_linear f | 
| 17 | using assms . | |
| 18 | show ?thesis | |
| 53600 
8fda7ad57466
make 'linear' into a sublocale of 'bounded_linear';
 huffman parents: 
53374diff
changeset | 19 | by (simp add: f.add f.scaleR linear_iff) | 
| 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 | 20 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 21 | |
| 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 | 22 | lemma netlimit_at_vector: (* TODO: move *) | 
| 37730 | 23 | fixes a :: "'a::real_normed_vector" | 
| 24 | shows "netlimit (at a) = a" | |
| 25 | proof (cases "\<exists>x. x \<noteq> a") | |
| 26 | case True then obtain x where x: "x \<noteq> a" .. | |
| 27 | have "\<not> trivial_limit (at a)" | |
| 28 | unfolding trivial_limit_def eventually_at dist_norm | |
| 29 | apply clarsimp | |
| 30 | apply (rule_tac x="a + scaleR (d / 2) (sgn (x - a))" in exI) | |
| 31 | apply (simp add: norm_sgn sgn_zero_iff x) | |
| 32 | done | |
| 53781 | 33 | then show ?thesis | 
| 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 | 34 | by (rule netlimit_within [of a UNIV]) | 
| 37730 | 35 | qed simp | 
| 36 | ||
| 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 | 37 | (* Because I do not want to type this all the time *) | 
| 53781 | 38 | lemmas linear_linear = linear_conv_bounded_linear[symmetric] | 
| 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 | 39 | |
| 53781 | 40 | lemma derivative_linear[dest]: "(f has_derivative f') net \<Longrightarrow> bounded_linear f'" | 
| 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 | 41 | unfolding has_derivative_def by auto | 
| 
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 | 42 | |
| 53781 | 43 | lemma derivative_is_linear: "(f has_derivative f') net \<Longrightarrow> linear f'" | 
| 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 | 44 | by (rule derivative_linear [THEN bounded_linear_imp_linear]) | 
| 44137 | 45 | |
| 53781 | 46 | lemma DERIV_conv_has_derivative: "(DERIV f x :> f') \<longleftrightarrow> (f has_derivative op * f') (at x)" | 
| 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 | 47 | using deriv_fderiv[of f x UNIV f'] by (subst (asm) mult_commute) | 
| 
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 | |
| 53781 | 49 | |
| 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 | 50 | subsection {* Derivatives *}
 | 
| 
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 | 51 | |
| 
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 | subsubsection {* Combining theorems. *}
 | 
| 
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 | 53 | |
| 
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 | 54 | lemmas has_derivative_id = FDERIV_ident | 
| 
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 | lemmas has_derivative_const = FDERIV_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 | 56 | lemmas has_derivative_neg = FDERIV_minus | 
| 
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 | lemmas has_derivative_add = FDERIV_add | 
| 
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 | 58 | lemmas has_derivative_sub = FDERIV_diff | 
| 
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 | 59 | lemmas has_derivative_setsum = FDERIV_setsum | 
| 
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 | lemmas scaleR_right_has_derivative = FDERIV_scaleR_right | 
| 
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 | lemmas scaleR_left_has_derivative = FDERIV_scaleR_left | 
| 
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 | lemmas inner_right_has_derivative = FDERIV_inner_right | 
| 
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 | lemmas inner_left_has_derivative = FDERIV_inner_left | 
| 
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 | lemmas mult_right_has_derivative = FDERIV_mult_right | 
| 
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 | 65 | lemmas mult_left_has_derivative = FDERIV_mult_left | 
| 
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 | 66 | |
| 
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 | 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 | 68 | "(f has_derivative f') net \<Longrightarrow> ((\<lambda>x. f x + c) has_derivative f') net" | 
| 
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 | 69 | by (intro FDERIV_eq_intros) auto | 
| 
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 | 70 | |
| 53781 | 71 | |
| 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 | 72 | subsection {* Derivative with composed bilinear function. *}
 | 
| 
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 | 73 | |
| 
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 | 74 | 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 | 75 | assumes "(f has_derivative f') (at x within s)" | 
| 53781 | 76 | and "(g has_derivative g') (at x within s)" | 
| 77 | 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 | 78 | 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 | 79 | 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 | 80 | |
| 
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 | 81 | 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 | 82 | assumes "(f has_derivative f') (at x)" | 
| 53781 | 83 | and "(g has_derivative g') (at x)" | 
| 84 | 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 | 85 | 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 | 86 | 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 | 87 | |
| 44124 
4c2a61a897d8
Derivative.thy: more sensible subsection headings
 huffman parents: 
44123diff
changeset | 88 | text {* These are the only cases we'll care about, probably. *}
 | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 89 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 90 | lemma has_derivative_within: "(f has_derivative f') (at x within s) \<longleftrightarrow> | 
| 53781 | 91 | bounded_linear f' \<and> ((\<lambda>y. (1 / norm(y - x)) *\<^sub>R (f y - (f x + f' (y - x)))) ---> 0) (at x within s)" | 
| 92 | unfolding has_derivative_def Lim | |
| 93 | by (auto simp add: netlimit_within inverse_eq_divide 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 | 94 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 95 | lemma has_derivative_at: "(f has_derivative f') (at x) \<longleftrightarrow> | 
| 53781 | 96 | bounded_linear f' \<and> ((\<lambda>y. (1 / (norm(y - x))) *\<^sub>R (f y - (f x + f' (y - x)))) ---> 0) (at x)" | 
| 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 | |
| 44124 
4c2a61a897d8
Derivative.thy: more sensible subsection headings
 huffman parents: 
44123diff
changeset | 100 | text {* More explicit epsilon-delta forms. *}
 | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 101 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 102 | lemma has_derivative_within': | 
| 53781 | 103 | "(f has_derivative f')(at x within s) \<longleftrightarrow> | 
| 104 | bounded_linear f' \<and> | |
| 105 | (\<forall>e>0. \<exists>d>0. \<forall>x'\<in>s. 0 < norm (x' - x) \<and> norm (x' - x) < d \<longrightarrow> | |
| 106 | norm (f x' - f x - f'(x' - x)) / norm (x' - x) < e)" | |
| 36587 | 107 | unfolding has_derivative_within Lim_within dist_norm | 
| 53781 | 108 | unfolding diff_0_right | 
| 109 | 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 | 110 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 111 | lemma has_derivative_at': | 
| 53781 | 112 | "(f has_derivative f') (at x) \<longleftrightarrow> bounded_linear f' \<and> | 
| 113 | (\<forall>e>0. \<exists>d>0. \<forall>x'. 0 < norm (x' - x) \<and> norm (x' - x) < d \<longrightarrow> | |
| 114 | norm (f x' - f x - f'(x' - x)) / norm (x' - x) < e)" | |
| 115 | using has_derivative_within' [of f f' x UNIV] | |
| 116 | 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 | 117 | |
| 53781 | 118 | lemma has_derivative_at_within: | 
| 119 | "(f has_derivative f') (at x) \<Longrightarrow> (f has_derivative f') (at x within s)" | |
| 120 | unfolding has_derivative_within' has_derivative_at' | |
| 121 | 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 | 122 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 123 | lemma has_derivative_within_open: | 
| 53781 | 124 | "a \<in> s \<Longrightarrow> open s \<Longrightarrow> | 
| 125 | (f has_derivative f') (at a within s) \<longleftrightarrow> (f has_derivative f') (at a)" | |
| 37730 | 126 | 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 | 127 | |
| 43338 | 128 | lemma has_derivative_right: | 
| 53781 | 129 | fixes f :: "real \<Rightarrow> real" | 
| 130 | and y :: "real" | |
| 43338 | 131 |   shows "(f has_derivative (op * y)) (at x within ({x <..} \<inter> I)) \<longleftrightarrow>
 | 
| 132 |     ((\<lambda>t. (f x - f t) / (x - t)) ---> y) (at x within ({x <..} \<inter> I))"
 | |
| 133 | proof - | |
| 134 |   have "((\<lambda>t. (f t - (f x + y * (t - x))) / \<bar>t - x\<bar>) ---> 0) (at x within ({x<..} \<inter> I)) \<longleftrightarrow>
 | |
| 135 |     ((\<lambda>t. (f t - f x) / (t - x) - y) ---> 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 | 136 | by (intro Lim_cong_within) (auto simp add: diff_divide_distrib add_divide_distrib) | 
| 43338 | 137 |   also have "\<dots> \<longleftrightarrow> ((\<lambda>t. (f t - f x) / (t - x)) ---> y) (at x within ({x<..} \<inter> I))"
 | 
| 138 | by (simp add: Lim_null[symmetric]) | |
| 139 |   also have "\<dots> \<longleftrightarrow> ((\<lambda>t. (f x - f t) / (x - t)) ---> y) (at x within ({x<..} \<inter> I))"
 | |
| 44140 
2c10c35dd4be
remove several redundant and unused theorems about derivatives
 huffman parents: 
44137diff
changeset | 140 | by (intro Lim_cong_within) (simp_all add: field_simps) | 
| 43338 | 141 | finally show ?thesis | 
| 44282 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44140diff
changeset | 142 | by (simp add: bounded_linear_mult_right has_derivative_within) | 
| 43338 | 143 | qed | 
| 144 | ||
| 53781 | 145 | |
| 44124 
4c2a61a897d8
Derivative.thy: more sensible subsection headings
 huffman parents: 
44123diff
changeset | 146 | subsubsection {* Limit transformation for derivatives *}
 | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 147 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 148 | lemma has_derivative_transform_within: | 
| 53781 | 149 | assumes "0 < d" | 
| 150 | and "x \<in> s" | |
| 151 | and "\<forall>x'\<in>s. dist x' x < d \<longrightarrow> f x' = g x'" | |
| 152 | and "(f has_derivative f') (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 | 153 | shows "(g has_derivative f') (at x within s)" | 
| 53781 | 154 | using assms(4) | 
| 155 | unfolding has_derivative_within | |
| 156 | apply - | |
| 157 | apply (erule conjE) | |
| 158 | apply rule | |
| 159 | apply assumption | |
| 160 | apply (rule Lim_transform_within[OF assms(1)]) | |
| 161 | defer | |
| 162 | apply assumption | |
| 163 | apply rule | |
| 164 | apply rule | |
| 165 | apply (drule assms(3)[rule_format]) | |
| 166 | using assms(3)[rule_format, OF assms(2)] | |
| 167 | apply auto | |
| 168 | done | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 169 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 170 | lemma has_derivative_transform_at: | 
| 53781 | 171 | assumes "0 < d" | 
| 172 | and "\<forall>x'. dist x' x < d \<longrightarrow> f x' = g x'" | |
| 173 | and "(f has_derivative f') (at x)" | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 174 | shows "(g has_derivative f') (at x)" | 
| 53781 | 175 | using has_derivative_transform_within [of d x UNIV f g f'] assms | 
| 176 | 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 | 177 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 178 | lemma has_derivative_transform_within_open: | 
| 53781 | 179 | assumes "open s" | 
| 180 | and "x \<in> s" | |
| 181 | and "\<forall>y\<in>s. f y = g y" | |
| 182 | and "(f has_derivative f') (at x)" | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 183 | shows "(g has_derivative f') (at x)" | 
| 53781 | 184 | using assms(4) | 
| 185 | unfolding has_derivative_at | |
| 186 | apply - | |
| 187 | apply (erule conjE) | |
| 188 | apply rule | |
| 189 | apply assumption | |
| 190 | apply (rule Lim_transform_within_open[OF assms(1,2)]) | |
| 191 | defer | |
| 192 | apply assumption | |
| 193 | apply rule | |
| 194 | apply rule | |
| 195 | apply (drule assms(3)[rule_format]) | |
| 196 | using assms(3)[rule_format, OF assms(2)] | |
| 197 | apply auto | |
| 198 | done | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 199 | |
| 44124 
4c2a61a897d8
Derivative.thy: more sensible subsection headings
 huffman parents: 
44123diff
changeset | 200 | subsection {* Differentiability *}
 | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 201 | |
| 36362 
06475a1547cb
fix lots of looping simp calls and other warnings
 huffman parents: 
36334diff
changeset | 202 | no_notation Deriv.differentiable (infixl "differentiable" 60) | 
| 
06475a1547cb
fix lots of looping simp calls and other warnings
 huffman parents: 
36334diff
changeset | 203 | |
| 53781 | 204 | abbreviation | 
| 205 |   differentiable :: "('a::real_normed_vector \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> 'a filter \<Rightarrow> bool"
 | |
| 206 | (infixr "differentiable" 30) | |
| 207 | where "f differentiable net \<equiv> isDiff net f" | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 208 | |
| 53781 | 209 | definition | 
| 210 |   differentiable_on :: "('a::real_normed_vector \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> 'a set \<Rightarrow> bool"
 | |
| 211 | (infixr "differentiable'_on" 30) | |
| 212 | 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 | 213 | |
| 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 | 214 | lemmas differentiable_def = isDiff_def | 
| 
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 | 215 | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 216 | lemma differentiableI: "(f has_derivative f') net \<Longrightarrow> f differentiable net" | 
| 53781 | 217 | unfolding differentiable_def | 
| 218 | 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 | 219 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 220 | lemma differentiable_at_withinI: "f differentiable (at x) \<Longrightarrow> f differentiable (at x within s)" | 
| 53781 | 221 | unfolding differentiable_def | 
| 222 | using has_derivative_at_within | |
| 223 | 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 | 224 | |
| 44123 | 225 | lemma differentiable_within_open: (* TODO: delete *) | 
| 53781 | 226 | assumes "a \<in> s" | 
| 227 | and "open s" | |
| 228 | shows "f differentiable (at a within s) \<longleftrightarrow> f differentiable (at a)" | |
| 229 | using assms | |
| 230 | 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 | 231 | |
| 44123 | 232 | lemma differentiable_on_eq_differentiable_at: | 
| 53781 | 233 | "open s \<Longrightarrow> f differentiable_on s \<longleftrightarrow> (\<forall>x\<in>s. f differentiable at x)" | 
| 44123 | 234 | 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 | 235 | 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 | 236 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 237 | lemma differentiable_transform_within: | 
| 53781 | 238 | assumes "0 < d" | 
| 239 | and "x \<in> s" | |
| 240 | and "\<forall>x'\<in>s. dist x' x < d \<longrightarrow> f x' = g x'" | |
| 44123 | 241 | assumes "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 | 242 | shows "g differentiable (at x within s)" | 
| 53781 | 243 | using assms(4) | 
| 244 | unfolding differentiable_def | |
| 44123 | 245 | by (auto intro!: has_derivative_transform_within[OF assms(1-3)]) | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 246 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 247 | lemma differentiable_transform_at: | 
| 53781 | 248 | assumes "0 < d" | 
| 249 | and "\<forall>x'. dist x' x < d \<longrightarrow> f x' = g x'" | |
| 250 | and "f differentiable at x" | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 251 | shows "g differentiable at x" | 
| 53781 | 252 | using assms(3) | 
| 253 | unfolding differentiable_def | |
| 254 | using has_derivative_transform_at[OF assms(1-2)] | |
| 255 | 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 | 256 | |
| 53781 | 257 | |
| 258 | subsection {* Frechet derivative and Jacobian matrix *}
 | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 259 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 260 | 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 | 261 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 262 | lemma frechet_derivative_works: | 
| 53781 | 263 | "f differentiable net \<longleftrightarrow> (f has_derivative (frechet_derivative f net)) net" | 
| 264 | unfolding frechet_derivative_def differentiable_def | |
| 265 | 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 | 266 | |
| 53781 | 267 | lemma linear_frechet_derivative: "f differentiable net \<Longrightarrow> linear(frechet_derivative f net)" | 
| 44123 | 268 | unfolding frechet_derivative_works has_derivative_def | 
| 269 | by (auto intro: bounded_linear_imp_linear) | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 270 | |
| 53781 | 271 | |
| 44124 
4c2a61a897d8
Derivative.thy: more sensible subsection headings
 huffman parents: 
44123diff
changeset | 272 | subsection {* Differentiability implies continuity *}
 | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 273 | |
| 44123 | 274 | lemma Lim_mul_norm_within: | 
| 53781 | 275 | fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" | 
| 276 | shows "(f ---> 0) (at a within s) \<Longrightarrow> ((\<lambda>x. norm(x - a) *\<^sub>R f x) ---> 0) (at a within s)" | |
| 277 | unfolding Lim_within | |
| 278 | apply rule | |
| 279 | apply rule | |
| 280 | apply (erule_tac x=e in allE) | |
| 281 | apply (erule impE) | |
| 282 | apply assumption | |
| 283 | apply (erule exE) | |
| 284 | apply (erule conjE) | |
| 285 | apply (rule_tac x="min d 1" in exI) | |
| 286 | apply rule | |
| 287 | defer | |
| 288 | apply rule | |
| 289 | apply (erule_tac x=x in ballE) | |
| 290 | unfolding dist_norm diff_0_right | |
| 291 | apply (auto intro!: mult_strict_mono[of _ "1::real", unfolded mult_1_left]) | |
| 292 | done | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 293 | |
| 44123 | 294 | 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 | 295 | "f differentiable (at x within s) \<Longrightarrow> continuous (at x within s) f" | 
| 
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 | 296 | by (auto simp: differentiable_def intro: FDERIV_continuous) | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 297 | |
| 44123 | 298 | lemma differentiable_imp_continuous_on: | 
| 299 | "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 | 300 | 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 | 301 | 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 | 302 | |
| 
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_within_subset: | 
| 53781 | 304 | "(f has_derivative f') (at x within s) \<Longrightarrow> t \<subseteq> s \<Longrightarrow> | 
| 305 | (f has_derivative f') (at x within t)" | |
| 306 | unfolding has_derivative_within | |
| 307 | using tendsto_within_subset | |
| 308 | 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 | 309 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 310 | lemma differentiable_within_subset: | 
| 53781 | 311 | "f differentiable (at x within t) \<Longrightarrow> s \<subseteq> t \<Longrightarrow> | 
| 312 | f differentiable (at x within s)" | |
| 313 | unfolding differentiable_def | |
| 314 | using has_derivative_within_subset | |
| 315 | 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 | 316 | |
| 44123 | 317 | lemma differentiable_on_subset: | 
| 318 | "f differentiable_on t \<Longrightarrow> s \<subseteq> t \<Longrightarrow> f differentiable_on s" | |
| 53781 | 319 | unfolding differentiable_on_def | 
| 320 | using differentiable_within_subset | |
| 321 | 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 | 322 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 323 | lemma differentiable_on_empty: "f differentiable_on {}"
 | 
| 53781 | 324 | unfolding differentiable_on_def | 
| 325 | 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 | 326 | |
| 44124 
4c2a61a897d8
Derivative.thy: more sensible subsection headings
 huffman parents: 
44123diff
changeset | 327 | text {* Several results are easier using a "multiplied-out" variant.
 | 
| 
4c2a61a897d8
Derivative.thy: more sensible subsection headings
 huffman parents: 
44123diff
changeset | 328 | (I got this idea from Dieudonne's proof of the chain rule). *} | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 329 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 330 | lemma has_derivative_within_alt: | 
| 53781 | 331 | "(f has_derivative f') (at x within s) \<longleftrightarrow> bounded_linear f' \<and> | 
| 332 | (\<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))" | |
| 333 | (is "?lhs \<longleftrightarrow> ?rhs") | |
| 44123 | 334 | proof | 
| 53781 | 335 | assume ?lhs | 
| 336 | then show ?rhs | |
| 337 | unfolding has_derivative_within | |
| 338 | apply - | |
| 339 | apply (erule conjE) | |
| 340 | apply rule | |
| 341 | apply assumption | |
| 44123 | 342 | unfolding Lim_within | 
| 53781 | 343 | apply rule | 
| 344 | apply (erule_tac x=e in allE) | |
| 345 | apply rule | |
| 346 | apply (erule impE) | |
| 347 | apply assumption | |
| 348 | apply (erule exE) | |
| 349 | apply (rule_tac x=d in exI) | |
| 350 | apply (erule conjE) | |
| 351 | apply rule | |
| 352 | apply assumption | |
| 353 | apply rule | |
| 354 | apply rule | |
| 44123 | 355 | proof- | 
| 53781 | 356 | fix x y e d | 
| 357 | assume as: | |
| 358 | "0 < e" | |
| 359 | "0 < d" | |
| 360 | "norm (y - x) < d" | |
| 361 | "\<forall>xa\<in>s. 0 < dist xa x \<and> dist xa x < d \<longrightarrow> | |
| 362 | dist ((1 / norm (xa - x)) *\<^sub>R (f xa - (f x + f' (xa - x)))) 0 < e" | |
| 363 | "y \<in> s" | |
| 364 | "bounded_linear f'" | |
| 365 | then interpret bounded_linear f' | |
| 366 | by auto | |
| 367 | show "norm (f y - f x - f' (y - x)) \<le> e * norm (y - x)" | |
| 368 | proof (cases "y = x") | |
| 369 | case True | |
| 370 | then show ?thesis | |
| 371 | using `bounded_linear f'` by (auto simp add: zero) | |
| 44123 | 372 | next | 
| 53781 | 373 | case False | 
| 374 | then have "norm (f y - (f x + f' (y - x))) < e * norm (y - x)" | |
| 375 | using as(4)[rule_format, OF `y \<in> s`] | |
| 376 | unfolding dist_norm diff_0_right | |
| 377 | using as(3) | |
| 41958 | 378 | using pos_divide_less_eq[OF False[unfolded dist_nz], unfolded dist_norm] | 
| 379 | by (auto simp add: linear_0 linear_sub) | |
| 53781 | 380 | then show ?thesis | 
| 381 | by (auto simp add: algebra_simps) | |
| 44123 | 382 | qed | 
| 383 | qed | |
| 384 | next | |
| 53781 | 385 | assume ?rhs | 
| 386 | then show ?lhs | |
| 387 | unfolding has_derivative_within Lim_within | |
| 388 | apply - | |
| 389 | apply (erule conjE) | |
| 390 | apply rule | |
| 391 | apply assumption | |
| 392 | apply rule | |
| 393 | apply (erule_tac x="e/2" in allE) | |
| 394 | apply rule | |
| 395 | apply (erule impE) | |
| 396 | defer | |
| 397 | apply (erule exE) | |
| 398 | apply (rule_tac x=d in exI) | |
| 399 | apply (erule conjE) | |
| 400 | apply rule | |
| 401 | apply assumption | |
| 402 | apply rule | |
| 403 | apply rule | |
| 44123 | 404 | unfolding dist_norm diff_0_right norm_scaleR | 
| 53781 | 405 | apply (erule_tac x=xa in ballE) | 
| 406 | apply (erule impE) | |
| 407 | proof - | |
| 408 | fix e d y | |
| 409 | assume "bounded_linear f'" | |
| 410 | and "0 < e" | |
| 411 | and "0 < d" | |
| 412 | and "y \<in> s" | |
| 413 | and "0 < norm (y - x) \<and> norm (y - x) < d" | |
| 414 | and "norm (f y - f x - f' (y - x)) \<le> e / 2 * norm (y - x)" | |
| 415 | then show "\<bar>1 / norm (y - x)\<bar> * norm (f y - (f x + f' (y - x))) < e" | |
| 416 | apply (rule_tac le_less_trans[of _ "e/2"]) | |
| 417 | apply (auto intro!: mult_imp_div_pos_le simp add: algebra_simps) | |
| 418 | done | |
| 44123 | 419 | qed auto | 
| 420 | qed | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 421 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 422 | lemma has_derivative_at_alt: | 
| 53781 | 423 | "(f has_derivative f') (at x) \<longleftrightarrow> | 
| 424 | bounded_linear f' \<and> | |
| 425 | (\<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))" | |
| 426 | using has_derivative_within_alt[where s=UNIV] | |
| 427 | 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 | 428 | |
| 53781 | 429 | |
| 430 | subsection {* The chain rule *}
 | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 431 | |
| 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 | 432 | lemma diff_chain_within[FDERIV_intros]: | 
| 44123 | 433 | assumes "(f has_derivative f') (at x within s)" | 
| 53781 | 434 | and "(g has_derivative g') (at (f x) within (f ` s))" | 
| 435 | shows "((g \<circ> f) has_derivative (g' \<circ> f'))(at x within s)" | |
| 436 | using FDERIV_in_compose[OF assms] | |
| 437 | 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 | 438 | |
| 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 | 439 | lemma diff_chain_at[FDERIV_intros]: | 
| 53781 | 440 | "(f has_derivative f') (at x) \<Longrightarrow> | 
| 441 | (g has_derivative g') (at (f x)) \<Longrightarrow> ((g \<circ> f) has_derivative (g' \<circ> f')) (at x)" | |
| 442 | using FDERIV_compose[of f f' x UNIV g g'] | |
| 443 | 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 | 444 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 445 | |
| 53781 | 446 | subsection {* Composition rules stated just for differentiability *}
 | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 447 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 448 | lemma differentiable_chain_at: | 
| 53781 | 449 | "f differentiable (at x) \<Longrightarrow> | 
| 450 | g differentiable (at (f x)) \<Longrightarrow> (g \<circ> f) differentiable (at x)" | |
| 451 | unfolding differentiable_def | |
| 452 | 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 | 453 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 454 | lemma differentiable_chain_within: | 
| 53781 | 455 | "f differentiable (at x within s) \<Longrightarrow> | 
| 456 | g differentiable (at(f x) within (f ` s)) \<Longrightarrow> (g \<circ> f) differentiable (at x within s)" | |
| 457 | unfolding differentiable_def | |
| 458 | by (meson diff_chain_within) | |
| 459 | ||
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 460 | |
| 37730 | 461 | subsection {* Uniqueness of derivative *}
 | 
| 462 | ||
| 463 | text {*
 | |
| 464 | The general result is a bit messy because we need approachability of the | |
| 465 | limit point from any direction. But OK for nontrivial intervals etc. | |
| 466 | *} | |
| 51363 
d4d00c804645
changed has_derivative_intros into a named theorems collection
 hoelzl parents: 
50939diff
changeset | 467 | |
| 44123 | 468 | lemma frechet_derivative_unique_within: | 
| 469 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector" | |
| 470 | assumes "(f has_derivative f') (at x within s)" | |
| 53781 | 471 | and "(f has_derivative f'') (at x within s)" | 
| 472 | and "\<forall>i\<in>Basis. \<forall>e>0. \<exists>d. 0 < abs d \<and> abs d < e \<and> (x + d *\<^sub>R i) \<in> s" | |
| 44123 | 473 | shows "f' = f''" | 
| 53781 | 474 | proof - | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 475 | note as = assms(1,2)[unfolded has_derivative_def] | 
| 44123 | 476 | then interpret f': bounded_linear f' by auto | 
| 477 | from as interpret f'': bounded_linear f'' by auto | |
| 478 | have "x islimpt s" unfolding islimpt_approachable | |
| 53781 | 479 | proof (rule, rule) | 
| 480 | fix e :: real | |
| 481 | assume "e > 0" | |
| 482 | guess d using assms(3)[rule_format,OF SOME_Basis `e>0`] .. | |
| 483 | then show "\<exists>x'\<in>s. x' \<noteq> x \<and> dist x' x < e" | |
| 484 | apply (rule_tac x="x + d *\<^sub>R (SOME i. i \<in> Basis)" in bexI) | |
| 485 | unfolding dist_norm | |
| 486 | apply (auto simp: SOME_Basis nonzero_Basis) | |
| 487 | done | |
| 44123 | 488 | qed | 
| 53781 | 489 | then have *: "netlimit (at x within s) = x" | 
| 490 | apply - | |
| 491 | apply (rule netlimit_within) | |
| 492 | unfolding trivial_limit_within | |
| 493 | apply simp | |
| 494 | done | |
| 495 | show ?thesis | |
| 496 | apply (rule linear_eq_stdbasis) | |
| 44123 | 497 | unfolding linear_conv_bounded_linear | 
| 53781 | 498 | apply (rule as(1,2)[THEN conjunct1])+ | 
| 499 | proof (rule, rule ccontr) | |
| 500 | fix i :: 'a | |
| 501 | assume i: "i \<in> Basis" | |
| 502 | 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 | 503 | assume "f' i \<noteq> f'' i" | 
| 53781 | 504 | then have "e > 0" | 
| 505 | unfolding e_def by auto | |
| 44125 | 506 | guess d using tendsto_diff [OF as(1,2)[THEN conjunct2], unfolded * Lim_within,rule_format,OF `e>0`] .. note d=this | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36844diff
changeset | 507 | guess c using assms(3)[rule_format,OF i d[THEN conjunct1]] .. note c=this | 
| 53781 | 508 | have *: "norm (- ((1 / \<bar>c\<bar>) *\<^sub>R f' (c *\<^sub>R i)) + (1 / \<bar>c\<bar>) *\<^sub>R f'' (c *\<^sub>R i)) = | 
| 509 | norm ((1 / abs c) *\<^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 | 510 | unfolding scaleR_right_distrib by auto | 
| 53781 | 511 | also have "\<dots> = norm ((1 / abs c) *\<^sub>R (c *\<^sub>R (- (f' i) + f'' i)))" | 
| 44123 | 512 | unfolding f'.scaleR f''.scaleR | 
| 53781 | 513 | unfolding scaleR_right_distrib scaleR_minus_right | 
| 514 | by auto | |
| 515 | also have "\<dots> = e" | |
| 516 | unfolding e_def | |
| 517 | using c[THEN conjunct1] | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50418diff
changeset | 518 | using norm_minus_cancel[of "f' i - f'' i"] | 
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
53799diff
changeset | 519 | by auto | 
| 53781 | 520 | finally show False | 
| 521 | using c | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50418diff
changeset | 522 | using d[THEN conjunct2,rule_format,of "x + c *\<^sub>R i"] | 
| 44123 | 523 | unfolding dist_norm | 
| 524 | unfolding f'.scaleR f''.scaleR f'.add f''.add f'.diff f''.diff | |
| 525 | scaleR_scaleR scaleR_right_diff_distrib scaleR_right_distrib | |
| 53781 | 526 | using i | 
| 527 | by (auto simp: inverse_eq_divide) | |
| 44123 | 528 | qed | 
| 529 | qed | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 530 | |
| 37730 | 531 | lemma frechet_derivative_unique_at: | 
| 53781 | 532 | "(f has_derivative f') (at x) \<Longrightarrow> (f has_derivative f'') (at x) \<Longrightarrow> f' = f''" | 
| 37730 | 533 | by (rule FDERIV_unique) | 
| 41829 | 534 | |
| 44123 | 535 | lemma frechet_derivative_unique_within_closed_interval: | 
| 536 | fixes f::"'a::ordered_euclidean_space \<Rightarrow> 'b::real_normed_vector" | |
| 53781 | 537 | assumes "\<forall>i\<in>Basis. a\<bullet>i < b\<bullet>i" | 
| 538 |     and "x \<in> {a..b}"
 | |
| 539 |     and "(f has_derivative f' ) (at x within {a..b})"
 | |
| 540 |     and "(f has_derivative f'') (at x within {a..b})"
 | |
| 44123 | 541 | shows "f' = f''" | 
| 542 | apply(rule frechet_derivative_unique_within) | |
| 543 | apply(rule assms(3,4))+ | |
| 53781 | 544 | proof (rule, rule, rule) | 
| 545 | fix e :: real | |
| 546 | fix i :: 'a | |
| 547 | assume "e > 0" and i: "i \<in> Basis" | |
| 548 |   then show "\<exists>d. 0 < \<bar>d\<bar> \<and> \<bar>d\<bar> < e \<and> x + d *\<^sub>R i \<in> {a..b}"
 | |
| 549 | proof (cases "x\<bullet>i = a\<bullet>i") | |
| 550 | case True | |
| 551 | then show ?thesis | |
| 552 | apply (rule_tac x="(min (b\<bullet>i - a\<bullet>i) e) / 2" in exI) | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50418diff
changeset | 553 | using assms(1)[THEN bspec[where x=i]] and `e>0` and assms(2) | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50418diff
changeset | 554 | unfolding mem_interval | 
| 53781 | 555 | using i | 
| 556 | apply (auto simp add: field_simps inner_simps inner_Basis) | |
| 557 | done | |
| 558 | next | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50418diff
changeset | 559 | note * = assms(2)[unfolded mem_interval, THEN bspec, OF i] | 
| 53781 | 560 | case False | 
| 561 | moreover have "a \<bullet> i < x \<bullet> i" | |
| 562 | using False * by auto | |
| 44123 | 563 |     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 | 564 | 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 | 565 | by auto | 
| 53781 | 566 | also have "\<dots> = a\<bullet>i + x\<bullet>i" | 
| 567 | by auto | |
| 568 | also have "\<dots> \<le> 2 * (x\<bullet>i)" | |
| 569 | using * by auto | |
| 570 | finally have "a \<bullet> i * 2 + min (x \<bullet> i - a \<bullet> i) e \<le> x \<bullet> i * 2" | |
| 571 | by auto | |
| 44123 | 572 | } | 
| 53781 | 573 | moreover have "min (x \<bullet> i - a \<bullet> i) e \<ge> 0" | 
| 574 | using * and `e>0` by auto | |
| 575 | then have "x \<bullet> i * 2 \<le> b \<bullet> i * 2 + min (x \<bullet> i - a \<bullet> i) e" | |
| 576 | using * by auto | |
| 44123 | 577 | ultimately show ?thesis | 
| 53781 | 578 | apply (rule_tac x="- (min (x\<bullet>i - a\<bullet>i) e) / 2" in exI) | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50418diff
changeset | 579 | using assms(1)[THEN bspec, OF i] and `e>0` and assms(2) | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50418diff
changeset | 580 | unfolding mem_interval | 
| 53781 | 581 | using i | 
| 582 | apply (auto simp add: field_simps inner_simps inner_Basis) | |
| 583 | done | |
| 44123 | 584 | qed | 
| 585 | qed | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 586 | |
| 44123 | 587 | lemma frechet_derivative_unique_within_open_interval: | 
| 588 | fixes f::"'a::ordered_euclidean_space \<Rightarrow> 'b::real_normed_vector" | |
| 589 |   assumes "x \<in> {a<..<b}"
 | |
| 53781 | 590 |     and "(f has_derivative f' ) (at x within {a<..<b})"
 | 
| 591 |     and "(f has_derivative f'') (at x within {a<..<b})"
 | |
| 37650 | 592 | shows "f' = f''" | 
| 593 | proof - | |
| 594 |   from assms(1) have *: "at x within {a<..<b} = at x"
 | |
| 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 | 595 | by (metis at_within_interior interior_open open_interval) | 
| 37650 | 596 | from assms(2,3) [unfolded *] show "f' = f''" | 
| 597 | by (rule frechet_derivative_unique_at) | |
| 598 | qed | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 599 | |
| 37730 | 600 | lemma frechet_derivative_at: | 
| 53781 | 601 | "(f has_derivative f') (at x) \<Longrightarrow> f' = frechet_derivative f (at x)" | 
| 602 | apply (rule frechet_derivative_unique_at[of f]) | |
| 603 | apply assumption | |
| 604 | unfolding frechet_derivative_works[symmetric] | |
| 605 | using differentiable_def | |
| 606 | apply auto | |
| 607 | done | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 608 | |
| 44123 | 609 | lemma frechet_derivative_within_closed_interval: | 
| 53781 | 610 | fixes f :: "'a::ordered_euclidean_space \<Rightarrow> 'b::real_normed_vector" | 
| 611 | assumes "\<forall>i\<in>Basis. a\<bullet>i < b\<bullet>i" | |
| 612 |     and "x \<in> {a..b}"
 | |
| 613 |     and "(f has_derivative f') (at x within {a..b})"
 | |
| 614 |   shows "frechet_derivative f (at x within {a..b}) = f'"
 | |
| 615 | apply (rule frechet_derivative_unique_within_closed_interval[where f=f]) | |
| 616 | apply (rule assms(1,2))+ | |
| 617 | unfolding frechet_derivative_works[symmetric] | |
| 618 | unfolding differentiable_def | |
| 619 | using assms(3) | |
| 620 | apply auto | |
| 621 | done | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 622 | |
| 53781 | 623 | |
| 624 | subsection {* The traditional Rolle theorem in one dimension *}
 | |
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36844diff
changeset | 625 | |
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36844diff
changeset | 626 | lemma linear_componentwise: | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36844diff
changeset | 627 | fixes f:: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36844diff
changeset | 628 | assumes lf: "linear f" | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50418diff
changeset | 629 | shows "(f x) \<bullet> j = (\<Sum>i\<in>Basis. (x\<bullet>i) * (f i\<bullet>j))" (is "?lhs = ?rhs") | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36844diff
changeset | 630 | proof - | 
| 53781 | 631 | have fA: "finite Basis" | 
| 632 | by simp | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50418diff
changeset | 633 | have "?rhs = (\<Sum>i\<in>Basis. (x\<bullet>i) *\<^sub>R (f i))\<bullet>j" | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50418diff
changeset | 634 | by (simp add: inner_setsum_left) | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36844diff
changeset | 635 | then show ?thesis | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36844diff
changeset | 636 | unfolding linear_setsum_mul[OF lf fA, symmetric] | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50418diff
changeset | 637 | unfolding euclidean_representation .. | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36844diff
changeset | 638 | qed | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36844diff
changeset | 639 | |
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36844diff
changeset | 640 | text {* We do not introduce @{text jacobian}, which is defined on matrices, instead we use
 | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36844diff
changeset | 641 | the unfolding of it. *} | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 642 | |
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36844diff
changeset | 643 | lemma jacobian_works: | 
| 53781 | 644 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 645 | shows "f differentiable net \<longleftrightarrow> | |
| 646 | (f has_derivative (\<lambda>h. \<Sum>i\<in>Basis. | |
| 647 | (\<Sum>j\<in>Basis. frechet_derivative f net j \<bullet> i * (h \<bullet> j)) *\<^sub>R i)) net" | |
| 648 | (is "?differentiable \<longleftrightarrow> (f has_derivative (\<lambda>h. \<Sum>i\<in>Basis. ?SUM h i *\<^sub>R i)) net") | |
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36844diff
changeset | 649 | proof | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36844diff
changeset | 650 | assume *: ?differentiable | 
| 53781 | 651 |   {
 | 
| 652 | fix h i | |
| 653 | have "?SUM h i = frechet_derivative f net h \<bullet> i" | |
| 654 | using * | |
| 655 | by (auto intro!: setsum_cong simp: linear_componentwise[of _ h i] linear_frechet_derivative) | |
| 656 | } | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50418diff
changeset | 657 | with * show "(f has_derivative (\<lambda>h. \<Sum>i\<in>Basis. ?SUM h i *\<^sub>R i)) net" | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50418diff
changeset | 658 | by (simp add: frechet_derivative_works euclidean_representation) | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36844diff
changeset | 659 | qed (auto intro!: differentiableI) | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36844diff
changeset | 660 | |
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36844diff
changeset | 661 | lemma differential_zero_maxmin_component: | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36844diff
changeset | 662 | 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 | 663 | assumes k: "k \<in> Basis" | 
| 53781 | 664 | 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 | 665 | 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 | 666 | 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 | 667 | proof (rule ccontr) | 
| 53781 | 668 | assume "\<not> ?thesis" | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50418diff
changeset | 669 | then obtain j where j: "?D k \<bullet> j \<noteq> 0" "j \<in> Basis" | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50418diff
changeset | 670 | unfolding euclidean_eq_iff[of _ "0::'a"] by auto | 
| 53781 | 671 | then have *: "\<bar>?D k \<bullet> j\<bar> / 2 > 0" | 
| 672 | by auto | |
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36844diff
changeset | 673 | note as = diff[unfolded jacobian_works has_derivative_at_alt] | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36844diff
changeset | 674 | guess e' using as[THEN conjunct2, rule_format, OF *] .. note e' = this | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36844diff
changeset | 675 | guess d using real_lbound_gt_zero[OF ball(1) e'[THEN conjunct1]] .. note d = this | 
| 53781 | 676 |   {
 | 
| 677 | fix c | |
| 678 | assume "abs c \<le> d" | |
| 679 | then have *: "norm (x + c *\<^sub>R j - x) < e'" | |
| 680 | using norm_Basis[OF j(2)] d by auto | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50418diff
changeset | 681 | let ?v = "(\<Sum>i\<in>Basis. (\<Sum>l\<in>Basis. ?D i \<bullet> l * ((c *\<^sub>R j :: 'a) \<bullet> l)) *\<^sub>R i)" | 
| 53781 | 682 | have if_dist: "\<And> P a b c. a * (if P then b else c) = (if P then a * b else a * c)" | 
| 683 | by auto | |
| 684 | have "\<bar>(f (x + c *\<^sub>R j) - f x - ?v) \<bullet> k\<bar> \<le> norm (f (x + c *\<^sub>R j) - f x - ?v)" | |
| 685 | by (rule Basis_le_norm[OF k]) | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50418diff
changeset | 686 | also have "\<dots> \<le> \<bar>?D k \<bullet> j\<bar> / 2 * \<bar>c\<bar>" | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50418diff
changeset | 687 | using e'[THEN conjunct2, rule_format, OF *] and norm_Basis[OF j(2)] j | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50418diff
changeset | 688 | by simp | 
| 53781 | 689 | finally have "\<bar>(f (x + c *\<^sub>R j) - f x - ?v) \<bullet> k\<bar> \<le> \<bar>?D k \<bullet> j\<bar> / 2 * \<bar>c\<bar>" | 
| 690 | by simp | |
| 691 | then have "\<bar>f (x + c *\<^sub>R j) \<bullet> k - f x \<bullet> k - c * (?D k \<bullet> j)\<bar> \<le> \<bar>?D k \<bullet> j\<bar> / 2 * \<bar>c\<bar>" | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50418diff
changeset | 692 | using j k | 
| 53781 | 693 | by (simp add: inner_simps field_simps inner_Basis setsum_cases if_dist) | 
| 694 | } | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50418diff
changeset | 695 | note * = this | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50418diff
changeset | 696 | have "x + d *\<^sub>R j \<in> ball x e" "x - d *\<^sub>R j \<in> ball x e" | 
| 53781 | 697 | unfolding mem_ball dist_norm | 
| 698 | using norm_Basis[OF j(2)] d | |
| 699 | by auto | |
| 700 | then have **: "((f (x - d *\<^sub>R j))\<bullet>k \<le> (f x)\<bullet>k \<and> (f (x + d *\<^sub>R j))\<bullet>k \<le> (f x)\<bullet>k) \<or> | |
| 701 | ((f (x - d *\<^sub>R j))\<bullet>k \<ge> (f x)\<bullet>k \<and> (f (x + d *\<^sub>R j))\<bullet>k \<ge> (f x)\<bullet>k)" | |
| 702 | using ball by auto | |
| 703 | have ***: "\<And>y y1 y2 d dx :: real. y1 \<le> y \<and> y2 \<le> y \<or> y \<le> y1 \<and> y \<le> y2 \<Longrightarrow> | |
| 704 | d < abs dx \<Longrightarrow> abs (y1 - y - - dx) \<le> d \<Longrightarrow> abs (y2 - y - dx) \<le> d \<Longrightarrow> False" | |
| 705 | by arith | |
| 706 | show False | |
| 707 | apply (rule ***[OF **, where dx="d * (?D k \<bullet> j)" and d="\<bar>?D k \<bullet> j\<bar> / 2 * \<bar>d\<bar>"]) | |
| 44123 | 708 | using *[of "-d"] and *[of d] and d[THEN conjunct1] and j | 
| 709 | unfolding mult_minus_left | |
| 44282 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44140diff
changeset | 710 | unfolding abs_mult diff_minus_eq_add scaleR_minus_left | 
| 53781 | 711 | unfolding algebra_simps | 
| 712 | apply (auto intro: mult_pos_pos) | |
| 713 | done | |
| 34906 | 714 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 715 | |
| 44124 
4c2a61a897d8
Derivative.thy: more sensible subsection headings
 huffman parents: 
44123diff
changeset | 716 | text {* In particular if we have a mapping into @{typ "real"}. *}
 | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 717 | |
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36844diff
changeset | 718 | lemma differential_zero_maxmin: | 
| 53781 | 719 | fixes f::"'a::euclidean_space \<Rightarrow> real" | 
| 720 | assumes "x \<in> s" | |
| 721 | and "open s" | |
| 722 | and deriv: "(f has_derivative f') (at x)" | |
| 723 | and mono: "(\<forall>y\<in>s. f y \<le> f x) \<or> (\<forall>y\<in>s. f x \<le> f y)" | |
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36844diff
changeset | 724 | shows "f' = (\<lambda>v. 0)" | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36844diff
changeset | 725 | proof - | 
| 53781 | 726 | obtain e where e: "e > 0" "ball x e \<subseteq> s" | 
| 44123 | 727 | using `open s`[unfolded open_contains_ball] and `x \<in> s` by auto | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50418diff
changeset | 728 | with differential_zero_maxmin_component[where 'b=real, of 1 e x f] mono deriv | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50418diff
changeset | 729 | have "(\<Sum>j\<in>Basis. frechet_derivative f (at x) j *\<^sub>R j) = (0::'a)" | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50418diff
changeset | 730 | by (auto simp: Basis_real_def differentiable_def) | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36844diff
changeset | 731 | with frechet_derivative_at[OF deriv, symmetric] | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50418diff
changeset | 732 | have "\<forall>i\<in>Basis. f' i = 0" | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50418diff
changeset | 733 | by (simp add: euclidean_eq_iff[of _ "0::'a"] inner_setsum_left_Basis) | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50418diff
changeset | 734 | with derivative_is_linear[OF deriv, THEN linear_componentwise, of _ 1] | 
| 53781 | 735 | show ?thesis | 
| 736 | by (simp add: fun_eq_iff) | |
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36844diff
changeset | 737 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 738 | |
| 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 | 739 | lemma rolle: | 
| 53781 | 740 | fixes f :: "real \<Rightarrow> real" | 
| 741 | assumes "a < b" | |
| 742 | and "f a = f b" | |
| 743 |     and "continuous_on {a..b} f"
 | |
| 744 |     and "\<forall>x\<in>{a<..<b}. (f has_derivative f' x) (at x)"
 | |
| 44123 | 745 |   shows "\<exists>x\<in>{a<..<b}. f' x = (\<lambda>v. 0)"
 | 
| 53781 | 746 | proof - | 
| 747 |   have "\<exists>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)"
 | |
| 748 | proof - | |
| 749 |     have "(a + b) / 2 \<in> {a .. b}"
 | |
| 750 | using assms(1) by auto | |
| 751 |     then have *: "{a..b} \<noteq> {}"
 | |
| 752 | by auto | |
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36844diff
changeset | 753 | guess d using continuous_attains_sup[OF compact_interval * assms(3)] .. note d=this | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36844diff
changeset | 754 | guess c using continuous_attains_inf[OF compact_interval * assms(3)] .. note c=this | 
| 44123 | 755 | show ?thesis | 
| 53781 | 756 |     proof (cases "d \<in> {a<..<b} \<or> c \<in> {a<..<b}")
 | 
| 757 | case True | |
| 758 | then show ?thesis | |
| 759 | apply (erule_tac disjE) | |
| 760 | apply (rule_tac x=d in bexI) | |
| 761 | apply (rule_tac[3] x=c in bexI) | |
| 762 | using d c | |
| 763 | apply auto | |
| 764 | done | |
| 44123 | 765 | next | 
| 766 | def e \<equiv> "(a + b) /2" | |
| 53781 | 767 | case False | 
| 768 | then have "f d = f c" | |
| 769 | using d c assms(2) by auto | |
| 770 |       then have "\<And>x. x \<in> {a..b} \<Longrightarrow> f x = f d"
 | |
| 771 | using c d | |
| 772 | apply - | |
| 773 | apply (erule_tac x=x in ballE)+ | |
| 774 | apply auto | |
| 775 | done | |
| 776 | then show ?thesis | |
| 777 | apply (rule_tac x=e in bexI) | |
| 778 | unfolding e_def | |
| 779 | using assms(1) | |
| 780 | apply auto | |
| 781 | done | |
| 44123 | 782 | qed | 
| 783 | qed | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 784 | then guess x .. note x=this | 
| 53781 | 785 | then have "f' x = (\<lambda>v. 0)" | 
| 786 |     apply (rule_tac differential_zero_maxmin[of x "{a<..<b}" f "f' x"])
 | |
| 787 | defer | |
| 788 | apply (rule open_interval) | |
| 789 | apply (rule assms(4)[unfolded has_derivative_at[symmetric],THEN bspec[where x=x]]) | |
| 790 | apply assumption | |
| 791 | unfolding o_def | |
| 792 | apply (erule disjE) | |
| 793 | apply (rule disjI2) | |
| 794 | apply auto | |
| 795 | done | |
| 796 | then show ?thesis | |
| 797 | apply (rule_tac x=x in bexI) | |
| 798 | unfolding o_def | |
| 799 | apply rule | |
| 800 | apply (drule_tac x=v in fun_cong) | |
| 801 | using x(1) | |
| 802 | apply auto | |
| 803 | done | |
| 44123 | 804 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 805 | |
| 53781 | 806 | |
| 807 | subsection {* One-dimensional mean value theorem *}
 | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 808 | |
| 53781 | 809 | lemma mvt: | 
| 810 | fixes f :: "real \<Rightarrow> real" | |
| 811 | assumes "a < b" | |
| 812 |     and "continuous_on {a..b} f"
 | |
| 44123 | 813 |   assumes "\<forall>x\<in>{a<..<b}. (f has_derivative (f' x)) (at x)"
 | 
| 53781 | 814 |   shows "\<exists>x\<in>{a<..<b}. f b - f a = (f' x) (b - a)"
 | 
| 815 | proof - | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 816 |   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 | 817 | proof (intro rolle[OF assms(1), of "\<lambda>x. f x - (f b - f a) / (b - a) * x"] ballI) | 
| 53781 | 818 | fix x | 
| 819 |     assume x: "x \<in> {a<..<b}"
 | |
| 820 | show "((\<lambda>x. f x - (f b - f a) / (b - a) * x) has_derivative | |
| 821 | (\<lambda>xa. f' x xa - (f b - f a) / (b - a) * xa)) (at x)" | |
| 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 | 822 | by (intro FDERIV_intros assms(3)[rule_format,OF x] mult_right_has_derivative) | 
| 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 | 823 | qed (insert assms(1,2), auto intro!: continuous_on_intros simp: field_simps) | 
| 44123 | 824 | then guess x .. | 
| 53781 | 825 | then show ?thesis | 
| 826 | apply (rule_tac x=x in bexI) | |
| 827 | apply (drule fun_cong[of _ _ "b - a"]) | |
| 828 | apply auto | |
| 829 | done | |
| 44123 | 830 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 831 | |
| 44123 | 832 | lemma mvt_simple: | 
| 53781 | 833 | fixes f :: "real \<Rightarrow> real" | 
| 834 | assumes "a < b" | |
| 835 |     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 | 836 |   shows "\<exists>x\<in>{a<..<b}. f b - f a = f' x (b - a)"
 | 
| 53781 | 837 | apply (rule mvt) | 
| 838 | apply (rule assms(1)) | |
| 839 | apply (rule differentiable_imp_continuous_on) | |
| 840 | unfolding differentiable_on_def differentiable_def | |
| 841 | defer | |
| 44123 | 842 | proof | 
| 53781 | 843 | fix x | 
| 844 |   assume x: "x \<in> {a<..<b}"
 | |
| 845 | show "(f has_derivative f' x) (at x)" | |
| 846 | unfolding has_derivative_within_open[OF x open_interval,symmetric] | |
| 847 | apply (rule has_derivative_within_subset) | |
| 848 | apply (rule assms(2)[rule_format]) | |
| 849 | using x | |
| 850 | apply auto | |
| 851 | done | |
| 852 | qed (insert assms(2), auto) | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 853 | |
| 44123 | 854 | lemma mvt_very_simple: | 
| 53781 | 855 | fixes f :: "real \<Rightarrow> real" | 
| 856 | assumes "a \<le> b" | |
| 857 |     and "\<forall>x\<in>{a..b}. (f has_derivative f' x) (at x within {a..b})"
 | |
| 44123 | 858 |   shows "\<exists>x\<in>{a..b}. f b - f a = f' x (b - a)"
 | 
| 859 | proof (cases "a = b") | |
| 53781 | 860 | interpret bounded_linear "f' b" | 
| 861 | using assms(2) assms(1) by auto | |
| 862 | case True | |
| 863 | then show ?thesis | |
| 864 | apply (rule_tac x=a in bexI) | |
| 865 | using assms(2)[THEN bspec[where x=a]] | |
| 866 | unfolding has_derivative_def | |
| 867 | unfolding True | |
| 868 | using zero | |
| 869 | apply auto | |
| 870 | done | |
| 871 | next | |
| 872 | case False | |
| 873 | then show ?thesis | |
| 874 | using mvt_simple[OF _ assms(2)] | |
| 875 | using assms(1) | |
| 876 | by auto | |
| 44123 | 877 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 878 | |
| 44124 
4c2a61a897d8
Derivative.thy: more sensible subsection headings
 huffman parents: 
44123diff
changeset | 879 | text {* A nice generalization (see Havin's proof of 5.19 from Rudin's book). *}
 | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 880 | |
| 44123 | 881 | lemma mvt_general: | 
| 53781 | 882 | fixes f :: "real \<Rightarrow> 'a::euclidean_space" | 
| 883 | assumes "a < b" | |
| 884 |     and "continuous_on {a..b} f"
 | |
| 885 |     and "\<forall>x\<in>{a<..<b}. (f has_derivative f'(x)) (at x)"
 | |
| 886 |   shows "\<exists>x\<in>{a<..<b}. norm (f b - f a) \<le> norm (f' x (b - a))"
 | |
| 887 | proof - | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 888 |   have "\<exists>x\<in>{a<..<b}. (op \<bullet> (f b - f a) \<circ> f) b - (op \<bullet> (f b - f a) \<circ> f) a = (f b - f a) \<bullet> f' x (b - a)"
 | 
| 53781 | 889 | apply (rule mvt) | 
| 890 | apply (rule assms(1)) | |
| 891 | apply (rule continuous_on_inner continuous_on_intros assms(2) ballI)+ | |
| 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 | 892 | unfolding o_def | 
| 53781 | 893 | apply (rule FDERIV_inner_right) | 
| 894 | using assms(3) | |
| 895 | apply auto | |
| 896 | done | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 897 | then guess x .. note x=this | 
| 53781 | 898 | show ?thesis | 
| 899 | proof (cases "f a = f b") | |
| 36844 | 900 | case False | 
| 53077 | 901 | have "norm (f b - f a) * norm (f b - f a) = (norm (f b - f a))\<^sup>2" | 
| 44123 | 902 | by (simp add: power2_eq_square) | 
| 53781 | 903 | also have "\<dots> = (f b - f a) \<bullet> (f b - f a)" | 
| 904 | unfolding power2_norm_eq_inner .. | |
| 44123 | 905 | also have "\<dots> = (f b - f a) \<bullet> f' x (b - a)" | 
| 53781 | 906 | using x | 
| 907 | unfolding inner_simps | |
| 908 | by (auto simp add: inner_diff_left) | |
| 44123 | 909 | also have "\<dots> \<le> norm (f b - f a) * norm (f' x (b - a))" | 
| 910 | by (rule norm_cauchy_schwarz) | |
| 53781 | 911 | finally show ?thesis | 
| 912 | using False x(1) | |
| 44123 | 913 | by (auto simp add: real_mult_left_cancel) | 
| 914 | next | |
| 53781 | 915 | case True | 
| 916 | then show ?thesis | |
| 917 | using assms(1) | |
| 918 | apply (rule_tac x="(a + b) /2" in bexI) | |
| 919 | apply auto | |
| 920 | done | |
| 44123 | 921 | qed | 
| 922 | qed | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 923 | |
| 44124 
4c2a61a897d8
Derivative.thy: more sensible subsection headings
 huffman parents: 
44123diff
changeset | 924 | text {* Still more general bound theorem. *}
 | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 925 | |
| 44123 | 926 | lemma differentiable_bound: | 
| 53781 | 927 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 928 | assumes "convex s" | |
| 929 | and "\<forall>x\<in>s. (f has_derivative f' x) (at x within s)" | |
| 930 | and "\<forall>x\<in>s. onorm (f' x) \<le> B" | |
| 931 | and x: "x \<in> s" | |
| 932 | and y: "y \<in> s" | |
| 933 | shows "norm (f x - f y) \<le> B * norm (x - y)" | |
| 934 | proof - | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 935 | let ?p = "\<lambda>u. x + u *\<^sub>R (y - x)" | 
| 53781 | 936 |   have *: "\<And>u. u\<in>{0..1} \<Longrightarrow> x + u *\<^sub>R (y - x) \<in> s"
 | 
| 44123 | 937 | using assms(1)[unfolded convex_alt,rule_format,OF x y] | 
| 938 | unfolding scaleR_left_diff_distrib scaleR_right_diff_distrib | |
| 939 | by (auto simp add: algebra_simps) | |
| 53781 | 940 |   then have 1: "continuous_on {0..1} (f \<circ> ?p)"
 | 
| 941 | apply - | |
| 942 | apply (rule continuous_on_intros)+ | |
| 44123 | 943 | unfolding continuous_on_eq_continuous_within | 
| 53781 | 944 | apply rule | 
| 945 | apply (rule differentiable_imp_continuous_within) | |
| 946 | unfolding differentiable_def | |
| 947 | apply (rule_tac x="f' xa" in exI) | |
| 948 | apply (rule has_derivative_within_subset) | |
| 949 | apply (rule assms(2)[rule_format]) | |
| 950 | apply auto | |
| 951 | done | |
| 952 |   have 2: "\<forall>u\<in>{0<..<1}.
 | |
| 953 | ((f \<circ> ?p) has_derivative f' (x + u *\<^sub>R (y - x)) \<circ> (\<lambda>u. 0 + u *\<^sub>R (y - x))) (at u)" | |
| 44123 | 954 | proof rule | 
| 955 | case goal1 | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 956 | let ?u = "x + u *\<^sub>R (y - x)" | 
| 53781 | 957 |     have "(f \<circ> ?p has_derivative (f' ?u) \<circ> (\<lambda>u. 0 + u *\<^sub>R (y - x))) (at u within {0<..<1})"
 | 
| 958 | apply (rule diff_chain_within) | |
| 959 | apply (rule FDERIV_intros)+ | |
| 960 | apply (rule has_derivative_within_subset) | |
| 961 | apply (rule assms(2)[rule_format]) | |
| 962 | using goal1 * | |
| 963 | apply auto | |
| 964 | done | |
| 965 | then show ?case | |
| 966 | unfolding has_derivative_within_open[OF goal1 open_interval] | |
| 967 | by auto | |
| 44123 | 968 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 969 | guess u using mvt_general[OF zero_less_one 1 2] .. note u = this | 
| 53781 | 970 | have **: "\<And>x y. x \<in> s \<Longrightarrow> norm (f' x y) \<le> B * norm y" | 
| 971 | proof - | |
| 44123 | 972 | case goal1 | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 973 | have "norm (f' x y) \<le> onorm (f' x) * norm y" | 
| 53781 | 974 | by (rule onorm(1)[OF derivative_is_linear[OF assms(2)[rule_format,OF goal1]]]) | 
| 44123 | 975 | also have "\<dots> \<le> B * norm y" | 
| 53781 | 976 | apply (rule mult_right_mono) | 
| 44123 | 977 | using assms(3)[rule_format,OF goal1] | 
| 53781 | 978 | apply (auto simp add: field_simps) | 
| 979 | done | |
| 980 | finally show ?case | |
| 981 | by simp | |
| 44123 | 982 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 983 | 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 | 984 | by (auto simp add: norm_minus_commute) | 
| 985 | also have "\<dots> \<le> norm (f' (x + u *\<^sub>R (y - x)) (y - x))" | |
| 986 | using u by auto | |
| 987 | also have "\<dots> \<le> B * norm(y - x)" | |
| 988 | apply (rule **) | |
| 989 | using * and u | |
| 990 | apply auto | |
| 991 | done | |
| 992 | finally show ?thesis | |
| 993 | by (auto simp add: norm_minus_commute) | |
| 44123 | 994 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 995 | |
| 44123 | 996 | lemma differentiable_bound_real: | 
| 53781 | 997 | fixes f :: "real \<Rightarrow> real" | 
| 998 | assumes "convex s" | |
| 999 | and "\<forall>x\<in>s. (f has_derivative f' x) (at x within s)" | |
| 1000 | and "\<forall>x\<in>s. onorm (f' x) \<le> B" | |
| 1001 | and x: "x \<in> s" | |
| 1002 | and y: "y \<in> s" | |
| 1003 | shows "norm (f x - f y) \<le> B * norm (x - y)" | |
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36844diff
changeset | 1004 | using differentiable_bound[of s f f' B x y] | 
| 53781 | 1005 | unfolding Ball_def image_iff o_def | 
| 1006 | using assms | |
| 1007 | by auto | |
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36844diff
changeset | 1008 | |
| 44124 
4c2a61a897d8
Derivative.thy: more sensible subsection headings
 huffman parents: 
44123diff
changeset | 1009 | text {* In particular. *}
 | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1010 | |
| 44123 | 1011 | lemma has_derivative_zero_constant: | 
| 53781 | 1012 | fixes f :: "real \<Rightarrow> real" | 
| 1013 | assumes "convex s" | |
| 1014 | and "\<forall>x\<in>s. (f has_derivative (\<lambda>h. 0)) (at x within s)" | |
| 44123 | 1015 | shows "\<exists>c. \<forall>x\<in>s. f x = c" | 
| 53781 | 1016 | proof (cases "s={}")
 | 
| 1017 | case False | |
| 1018 | then obtain x where "x \<in> s" | |
| 1019 | by auto | |
| 1020 | have "\<And>y. y \<in> s \<Longrightarrow> f x = f y" | |
| 1021 | proof - | |
| 1022 | case goal1 | |
| 1023 | then show ?case | |
| 1024 | using differentiable_bound_real[OF assms(1-2), of 0 x y] and `x \<in> s` | |
| 1025 | unfolding onorm_const | |
| 1026 | by auto | |
| 1027 | qed | |
| 1028 | then show ?thesis | |
| 1029 | apply (rule_tac x="f x" in exI) | |
| 1030 | apply auto | |
| 1031 | done | |
| 1032 | next | |
| 1033 | case True | |
| 1034 | then show ?thesis by auto | |
| 1035 | qed | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1036 | |
| 53781 | 1037 | lemma has_derivative_zero_unique: | 
| 1038 | fixes f :: "real \<Rightarrow> real" | |
| 1039 | assumes "convex s" | |
| 1040 | and "a \<in> s" | |
| 1041 | and "f a = c" | |
| 1042 | and "\<forall>x\<in>s. (f has_derivative (\<lambda>h. 0)) (at x within s)" | |
| 1043 | and "x \<in> s" | |
| 44123 | 1044 | shows "f x = c" | 
| 53781 | 1045 | using has_derivative_zero_constant[OF assms(1,4)] | 
| 1046 | using assms(2-3,5) | |
| 1047 | 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 | 1048 | |
| 53781 | 1049 | |
| 1050 | subsection {* Differentiability of inverse function (most basic form) *}
 | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1051 | |
| 44123 | 1052 | lemma has_derivative_inverse_basic: | 
| 53781 | 1053 | fixes f :: "'b::euclidean_space \<Rightarrow> 'c::euclidean_space" | 
| 44123 | 1054 | assumes "(f has_derivative f') (at (g y))" | 
| 53781 | 1055 | and "bounded_linear g'" | 
| 1056 | and "g' \<circ> f' = id" | |
| 1057 | and "continuous (at y) g" | |
| 1058 | and "open t" | |
| 1059 | and "y \<in> t" | |
| 1060 | and "\<forall>z\<in>t. f (g z) = z" | |
| 44123 | 1061 | shows "(g has_derivative g') (at y)" | 
| 53781 | 1062 | proof - | 
| 44123 | 1063 | interpret f': bounded_linear f' | 
| 1064 | using assms unfolding has_derivative_def by auto | |
| 53781 | 1065 | interpret g': bounded_linear g' | 
| 1066 | 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 | 1067 | guess C using bounded_linear.pos_bounded[OF assms(2)] .. note C = this | 
| 53781 | 1068 | have lem1: "\<forall>e>0. \<exists>d>0. \<forall>z. | 
| 1069 | norm (z - y) < d \<longrightarrow> norm (g z - g y - g'(z - y)) \<le> e * norm (g z - g y)" | |
| 1070 | proof (rule, rule) | |
| 44123 | 1071 | case goal1 | 
| 53781 | 1072 | have *: "e / C > 0" | 
| 1073 | apply (rule divide_pos_pos) | |
| 1074 | using `e > 0` C | |
| 1075 | apply auto | |
| 1076 | done | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1077 | guess d0 using assms(1)[unfolded has_derivative_at_alt,THEN conjunct2,rule_format,OF *] .. note d0=this | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1078 | guess d1 using assms(4)[unfolded continuous_at Lim_at,rule_format,OF d0[THEN conjunct1]] .. note d1=this | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1079 | guess d2 using assms(5)[unfolded open_dist,rule_format,OF assms(6)] .. note d2=this | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1080 | guess d using real_lbound_gt_zero[OF d1[THEN conjunct1] d2[THEN conjunct1]] .. note d=this | 
| 53781 | 1081 | then show ?case | 
| 1082 | apply (rule_tac x=d in exI) | |
| 1083 | apply rule | |
| 1084 | defer | |
| 1085 | apply rule | |
| 1086 | apply rule | |
| 1087 | proof - | |
| 1088 | fix z | |
| 1089 | assume as: "norm (z - y) < d" | |
| 1090 | then have "z \<in> t" | |
| 44123 | 1091 | 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 | 1092 | have "norm (g z - g y - g' (z - y)) \<le> norm (g' (f (g z) - y - f' (g z - g y)))" | 
| 44123 | 1093 | unfolding g'.diff f'.diff | 
| 53781 | 1094 | unfolding assms(3)[unfolded o_def id_def, THEN fun_cong] | 
| 44123 | 1095 | unfolding assms(7)[rule_format,OF `z\<in>t`] | 
| 53781 | 1096 | apply (subst norm_minus_cancel[symmetric]) | 
| 1097 | apply auto | |
| 1098 | done | |
| 1099 | also have "\<dots> \<le> norm (f (g z) - y - f' (g z - g y)) * C" | |
| 44123 | 1100 | by (rule C [THEN conjunct2, rule_format]) | 
| 1101 | also have "\<dots> \<le> (e / C) * norm (g z - g y) * C" | |
| 53781 | 1102 | apply (rule mult_right_mono) | 
| 1103 | apply (rule d0[THEN conjunct2,rule_format,unfolded assms(7)[rule_format,OF `y\<in>t`]]) | |
| 1104 | apply (cases "z = y") | |
| 1105 | defer | |
| 1106 | apply (rule d1[THEN conjunct2, unfolded dist_norm,rule_format]) | |
| 1107 | using as d C d0 | |
| 1108 | apply auto | |
| 1109 | done | |
| 44123 | 1110 | also have "\<dots> \<le> e * norm (g z - g y)" | 
| 1111 | using C by (auto simp add: field_simps) | |
| 1112 | finally show "norm (g z - g y - g' (z - y)) \<le> e * norm (g z - g y)" | |
| 1113 | by simp | |
| 1114 | qed auto | |
| 1115 | qed | |
| 53781 | 1116 | have *: "(0::real) < 1 / 2" | 
| 1117 | by auto | |
| 44123 | 1118 | guess d using lem1[rule_format,OF *] .. note d=this | 
| 53781 | 1119 | def B\<equiv>"C * 2" | 
| 1120 | have "B > 0" | |
| 1121 | unfolding B_def using C by auto | |
| 1122 | have lem2: "\<forall>z. norm(z - y) < d \<longrightarrow> norm (g z - g y) \<le> B * norm (z - y)" | |
| 1123 | proof (rule, rule) | |
| 1124 | case goal1 | |
| 44123 | 1125 | have "norm (g z - g y) \<le> norm(g' (z - y)) + norm ((g z - g y) - g'(z - y))" | 
| 53781 | 1126 | by (rule norm_triangle_sub) | 
| 1127 | also have "\<dots> \<le> norm (g' (z - y)) + 1 / 2 * norm (g z - g y)" | |
| 1128 | apply (rule add_left_mono) | |
| 1129 | using d and goal1 | |
| 1130 | apply auto | |
| 1131 | done | |
| 44123 | 1132 | also have "\<dots> \<le> norm (z - y) * C + 1 / 2 * norm (g z - g y)" | 
| 53781 | 1133 | apply (rule add_right_mono) | 
| 1134 | using C | |
| 1135 | apply auto | |
| 1136 | done | |
| 1137 | finally show ?case | |
| 1138 | unfolding B_def | |
| 1139 | by (auto simp add: field_simps) | |
| 44123 | 1140 | qed | 
| 53781 | 1141 | show ?thesis | 
| 1142 | unfolding has_derivative_at_alt | |
| 1143 | apply rule | |
| 1144 | apply (rule assms) | |
| 1145 | apply rule | |
| 1146 | apply rule | |
| 1147 | proof - | |
| 1148 | case goal1 | |
| 1149 | then have *: "e / B >0" | |
| 1150 | apply - | |
| 1151 | apply (rule divide_pos_pos) | |
| 1152 | using `B > 0` | |
| 1153 | apply auto | |
| 1154 | done | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1155 | guess d' using lem1[rule_format,OF *] .. note d'=this | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1156 | guess k using real_lbound_gt_zero[OF d[THEN conjunct1] d'[THEN conjunct1]] .. note k=this | 
| 44123 | 1157 | show ?case | 
| 53781 | 1158 | apply (rule_tac x=k in exI) | 
| 1159 | apply rule | |
| 1160 | defer | |
| 1161 | apply rule | |
| 1162 | apply rule | |
| 1163 | proof - | |
| 1164 | fix z | |
| 1165 | assume as: "norm (z - y) < k" | |
| 1166 | then have "norm (g z - g y - g' (z - y)) \<le> e / B * norm(g z - g y)" | |
| 44123 | 1167 | using d' k by auto | 
| 53781 | 1168 | also have "\<dots> \<le> e * norm (z - y)" | 
| 44123 | 1169 | unfolding times_divide_eq_left pos_divide_le_eq[OF `B>0`] | 
| 53781 | 1170 | using lem2[THEN spec[where x=z]] | 
| 1171 | using k as using `e > 0` | |
| 44123 | 1172 | by (auto simp add: field_simps) | 
| 1173 | finally show "norm (g z - g y - g' (z - y)) \<le> e * norm (z - y)" | |
| 53781 | 1174 | by simp | 
| 1175 | qed(insert k, auto) | |
| 44123 | 1176 | qed | 
| 1177 | qed | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1178 | |
| 44124 
4c2a61a897d8
Derivative.thy: more sensible subsection headings
 huffman parents: 
44123diff
changeset | 1179 | text {* Simply rewrite that based on the domain point x. *}
 | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1180 | |
| 44123 | 1181 | lemma has_derivative_inverse_basic_x: | 
| 53781 | 1182 | fixes f :: "'b::euclidean_space \<Rightarrow> 'c::euclidean_space" | 
| 1183 | assumes "(f has_derivative f') (at x)" | |
| 1184 | and "bounded_linear g'" | |
| 1185 | and "g' \<circ> f' = id" | |
| 1186 | and "continuous (at (f x)) g" | |
| 1187 | and "g (f x) = x" | |
| 1188 | and "open t" | |
| 1189 | and "f x \<in> t" | |
| 1190 | and "\<forall>y\<in>t. f (g y) = y" | |
| 1191 | shows "(g has_derivative g') (at (f x))" | |
| 1192 | apply (rule has_derivative_inverse_basic) | |
| 1193 | using assms | |
| 1194 | apply auto | |
| 1195 | done | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1196 | |
| 44124 
4c2a61a897d8
Derivative.thy: more sensible subsection headings
 huffman parents: 
44123diff
changeset | 1197 | text {* This is the version in Dieudonne', assuming continuity of f and g. *}
 | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1198 | |
| 44123 | 1199 | lemma has_derivative_inverse_dieudonne: | 
| 53781 | 1200 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 1201 | assumes "open s" | |
| 1202 | and "open (f ` s)" | |
| 1203 | and "continuous_on s f" | |
| 1204 | and "continuous_on (f ` s) g" | |
| 1205 | and "\<forall>x\<in>s. g (f x) = x" | |
| 1206 | and "x \<in> s" | |
| 1207 | and "(f has_derivative f') (at x)" | |
| 1208 | and "bounded_linear g'" | |
| 1209 | 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 | 1210 | shows "(g has_derivative g') (at (f x))" | 
| 53781 | 1211 | apply (rule has_derivative_inverse_basic_x[OF assms(7-9) _ _ assms(2)]) | 
| 1212 | using assms(3-6) | |
| 1213 | unfolding continuous_on_eq_continuous_at[OF assms(1)] continuous_on_eq_continuous_at[OF assms(2)] | |
| 1214 | apply auto | |
| 1215 | done | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1216 | |
| 44124 
4c2a61a897d8
Derivative.thy: more sensible subsection headings
 huffman parents: 
44123diff
changeset | 1217 | text {* Here's the simplest way of not assuming much about g. *}
 | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1218 | |
| 44123 | 1219 | lemma has_derivative_inverse: | 
| 53781 | 1220 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 1221 | assumes "compact s" | |
| 1222 | and "x \<in> s" | |
| 1223 | and "f x \<in> interior (f ` s)" | |
| 1224 | and "continuous_on s f" | |
| 1225 | and "\<forall>y\<in>s. g (f y) = y" | |
| 1226 | and "(f has_derivative f') (at x)" | |
| 1227 | and "bounded_linear g'" | |
| 1228 | and "g' \<circ> f' = id" | |
| 44123 | 1229 | shows "(g has_derivative g') (at (f x))" | 
| 53781 | 1230 | proof - | 
| 1231 |   {
 | |
| 1232 | fix y | |
| 1233 | assume "y \<in> interior (f ` s)" | |
| 1234 | then obtain x where "x \<in> s" and *: "y = f x" | |
| 1235 | unfolding image_iff | |
| 1236 | using interior_subset | |
| 1237 | by auto | |
| 1238 | have "f (g y) = y" | |
| 1239 | unfolding * and assms(5)[rule_format,OF `x\<in>s`] .. | |
| 44123 | 1240 | } note * = this | 
| 1241 | show ?thesis | |
| 53781 | 1242 | apply (rule has_derivative_inverse_basic_x[OF assms(6-8)]) | 
| 1243 | apply (rule continuous_on_interior[OF _ assms(3)]) | |
| 1244 | apply (rule continuous_on_inv[OF assms(4,1)]) | |
| 1245 | apply (rule assms(2,5) assms(5)[rule_format] open_interior assms(3))+ | |
| 1246 | apply rule | |
| 1247 | apply (rule *) | |
| 1248 | apply assumption | |
| 1249 | done | |
| 44123 | 1250 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1251 | |
| 53781 | 1252 | |
| 1253 | subsection {* Proving surjectivity via Brouwer fixpoint theorem *}
 | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1254 | |
| 44123 | 1255 | lemma brouwer_surjective: | 
| 53781 | 1256 | fixes f :: "'n::ordered_euclidean_space \<Rightarrow> 'n" | 
| 1257 | assumes "compact t" | |
| 1258 | and "convex t" | |
| 1259 |     and "t \<noteq> {}"
 | |
| 1260 | and "continuous_on t f" | |
| 1261 | and "\<forall>x\<in>s. \<forall>y\<in>t. x + (y - f y) \<in> t" | |
| 1262 | and "x \<in> s" | |
| 44123 | 1263 | shows "\<exists>y\<in>t. f y = x" | 
| 53781 | 1264 | proof - | 
| 1265 | have *: "\<And>x y. f y = x \<longleftrightarrow> x + (y - f y) = y" | |
| 1266 | by (auto simp add: algebra_simps) | |
| 44123 | 1267 | show ?thesis | 
| 1268 | unfolding * | |
| 53781 | 1269 | apply (rule brouwer[OF assms(1-3), of "\<lambda>y. x + (y - f y)"]) | 
| 1270 | apply (rule continuous_on_intros assms)+ | |
| 1271 | using assms(4-6) | |
| 1272 | apply auto | |
| 1273 | done | |
| 44123 | 1274 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1275 | |
| 44123 | 1276 | lemma brouwer_surjective_cball: | 
| 53781 | 1277 | fixes f :: "'n::ordered_euclidean_space \<Rightarrow> 'n" | 
| 1278 | assumes "e > 0" | |
| 1279 | and "continuous_on (cball a e) f" | |
| 1280 | and "\<forall>x\<in>s. \<forall>y\<in>cball a e. x + (y - f y) \<in> cball a e" | |
| 1281 | and "x \<in> s" | |
| 44123 | 1282 | shows "\<exists>y\<in>cball a e. f y = x" | 
| 53781 | 1283 | apply (rule brouwer_surjective) | 
| 1284 | apply (rule compact_cball convex_cball)+ | |
| 1285 | unfolding cball_eq_empty | |
| 1286 | using assms | |
| 1287 | apply auto | |
| 1288 | done | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1289 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1290 | text {* See Sussmann: "Multidifferential calculus", Theorem 2.1.1 *}
 | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1291 | |
| 44123 | 1292 | lemma sussmann_open_mapping: | 
| 53781 | 1293 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::ordered_euclidean_space" | 
| 1294 | assumes "open s" | |
| 1295 | and "continuous_on s f" | |
| 1296 | and "x \<in> s" | |
| 1297 | and "(f has_derivative f') (at x)" | |
| 1298 | and "bounded_linear g'" "f' \<circ> g' = id" | |
| 1299 | and "t \<subseteq> s" | |
| 1300 | and "x \<in> interior t" | |
| 44123 | 1301 | shows "f x \<in> interior (f ` t)" | 
| 53781 | 1302 | proof - | 
| 1303 | interpret f': bounded_linear f' | |
| 1304 | using assms | |
| 1305 | unfolding has_derivative_def | |
| 1306 | by auto | |
| 1307 | interpret g': bounded_linear g' | |
| 1308 | using assms | |
| 1309 | by auto | |
| 44123 | 1310 | guess B using bounded_linear.pos_bounded[OF assms(5)] .. note B=this | 
| 53781 | 1311 | then have *: "1 / (2 * B) > 0" | 
| 1312 | by (auto intro!: divide_pos_pos) | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1313 | guess e0 using assms(4)[unfolded has_derivative_at_alt,THEN conjunct2,rule_format,OF *] .. note e0=this | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1314 | guess e1 using assms(8)[unfolded mem_interior_cball] .. note e1=this | 
| 53781 | 1315 | have *: "0 < e0 / B" "0 < e1 / B" | 
| 1316 | apply (rule_tac[!] divide_pos_pos) | |
| 1317 | using e0 e1 B | |
| 1318 | apply auto | |
| 1319 | done | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1320 | guess e using real_lbound_gt_zero[OF *] .. note e=this | 
| 53781 | 1321 | have "\<forall>z\<in>cball (f x) (e / 2). \<exists>y\<in>cball (f x) e. f (x + g' (y - f x)) = z" | 
| 1322 | apply rule | |
| 1323 | apply (rule brouwer_surjective_cball[where s="cball (f x) (e/2)"]) | |
| 1324 | prefer 3 | |
| 1325 | apply rule | |
| 1326 | apply rule | |
| 44123 | 1327 | proof- | 
| 1328 | show "continuous_on (cball (f x) e) (\<lambda>y. f (x + g' (y - f x)))" | |
| 1329 | unfolding g'.diff | |
| 53781 | 1330 | apply (rule continuous_on_compose[of _ _ f, unfolded o_def]) | 
| 1331 | apply (rule continuous_on_intros linear_continuous_on[OF assms(5)])+ | |
| 1332 | apply (rule continuous_on_subset[OF assms(2)]) | |
| 1333 | apply rule | |
| 1334 | apply (unfold image_iff) | |
| 1335 | apply (erule bexE) | |
| 44123 | 1336 | proof- | 
| 53781 | 1337 | fix y z | 
| 1338 | assume as: "y \<in>cball (f x) e" "z = x + (g' y - g' (f x))" | |
| 44123 | 1339 | have "dist x z = norm (g' (f x) - g' y)" | 
| 1340 | unfolding as(2) and dist_norm by auto | |
| 1341 | also have "\<dots> \<le> norm (f x - y) * B" | |
| 53781 | 1342 | unfolding g'.diff[symmetric] | 
| 1343 | using B | |
| 1344 | by auto | |
| 44123 | 1345 | also have "\<dots> \<le> e * B" | 
| 53781 | 1346 | using as(1)[unfolded mem_cball dist_norm] | 
| 1347 | using B | |
| 1348 | by auto | |
| 1349 | also have "\<dots> \<le> e1" | |
| 1350 | using e | |
| 1351 | unfolding less_divide_eq | |
| 1352 | using B | |
| 1353 | by auto | |
| 1354 | finally have "z \<in> cball x e1" | |
| 1355 | unfolding mem_cball | |
| 1356 | by force | |
| 1357 | then show "z \<in> s" | |
| 1358 | using e1 assms(7) by auto | |
| 44123 | 1359 | qed | 
| 1360 | next | |
| 53781 | 1361 | fix y z | 
| 1362 | assume as: "y \<in> cball (f x) (e / 2)" "z \<in> cball (f x) e" | |
| 1363 | have "norm (g' (z - f x)) \<le> norm (z - f x) * B" | |
| 1364 | using B by auto | |
| 1365 | also have "\<dots> \<le> e * B" | |
| 1366 | apply (rule mult_right_mono) | |
| 44123 | 1367 | using as(2)[unfolded mem_cball dist_norm] and B | 
| 53781 | 1368 | unfolding norm_minus_commute | 
| 1369 | apply auto | |
| 1370 | done | |
| 1371 | also have "\<dots> < e0" | |
| 1372 | using e and B | |
| 1373 | unfolding less_divide_eq | |
| 1374 | by auto | |
| 1375 | finally have *: "norm (x + g' (z - f x) - x) < e0" | |
| 1376 | by auto | |
| 1377 | have **: "f x + f' (x + g' (z - f x) - x) = z" | |
| 1378 | using assms(6)[unfolded o_def id_def,THEN cong] | |
| 1379 | by auto | |
| 1380 | have "norm (f x - (y + (z - f (x + g' (z - f x))))) \<le> | |
| 1381 | norm (f (x + g' (z - f x)) - z) + norm (f x - y)" | |
| 44123 | 1382 | using norm_triangle_ineq[of "f (x + g'(z - f x)) - z" "f x - y"] | 
| 1383 | by (auto simp add: algebra_simps) | |
| 1384 | also have "\<dots> \<le> 1 / (B * 2) * norm (g' (z - f x)) + norm (f x - y)" | |
| 1385 | using e0[THEN conjunct2,rule_format,OF *] | |
| 53781 | 1386 | unfolding algebra_simps ** | 
| 1387 | by auto | |
| 44123 | 1388 | also have "\<dots> \<le> 1 / (B * 2) * norm (g' (z - f x)) + e/2" | 
| 53781 | 1389 | using as(1)[unfolded mem_cball dist_norm] | 
| 1390 | by auto | |
| 44123 | 1391 | also have "\<dots> \<le> 1 / (B * 2) * B * norm (z - f x) + e/2" | 
| 53781 | 1392 | using * and B | 
| 1393 | by (auto simp add: field_simps) | |
| 1394 | also have "\<dots> \<le> 1 / 2 * norm (z - f x) + e/2" | |
| 1395 | by auto | |
| 1396 | also have "\<dots> \<le> e/2 + e/2" | |
| 1397 | apply (rule add_right_mono) | |
| 44123 | 1398 | using as(2)[unfolded mem_cball dist_norm] | 
| 53781 | 1399 | unfolding norm_minus_commute | 
| 1400 | apply auto | |
| 1401 | done | |
| 44123 | 1402 | finally show "y + (z - f (x + g' (z - f x))) \<in> cball (f x) e" | 
| 53781 | 1403 | unfolding mem_cball dist_norm | 
| 1404 | by auto | |
| 1405 | qed (insert e, auto) note lem = this | |
| 1406 | show ?thesis | |
| 1407 | unfolding mem_interior | |
| 1408 | apply (rule_tac x="e/2" in exI) | |
| 1409 | apply rule | |
| 1410 | apply (rule divide_pos_pos) | |
| 1411 | prefer 3 | |
| 44123 | 1412 | proof | 
| 53781 | 1413 | fix y | 
| 1414 | assume "y \<in> ball (f x) (e / 2)" | |
| 1415 | then have *: "y \<in> cball (f x) (e / 2)" | |
| 1416 | 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 | 1417 | guess z using lem[rule_format,OF *] .. note z=this | 
| 53781 | 1418 | then have "norm (g' (z - f x)) \<le> norm (z - f x) * B" | 
| 1419 | using B | |
| 1420 | by (auto simp add: field_simps) | |
| 44123 | 1421 | also have "\<dots> \<le> e * B" | 
| 53781 | 1422 | apply (rule mult_right_mono) | 
| 1423 | using z(1) | |
| 1424 | unfolding mem_cball dist_norm norm_minus_commute | |
| 1425 | using B | |
| 1426 | apply auto | |
| 1427 | done | |
| 1428 | also have "\<dots> \<le> e1" | |
| 1429 | using e B unfolding less_divide_eq by auto | |
| 1430 | finally have "x + g'(z - f x) \<in> t" | |
| 1431 | apply - | |
| 1432 | apply (rule e1[THEN conjunct2,unfolded subset_eq,rule_format]) | |
| 1433 | unfolding mem_cball dist_norm | |
| 1434 | apply auto | |
| 1435 | done | |
| 1436 | then show "y \<in> f ` t" | |
| 1437 | using z by auto | |
| 1438 | qed (insert e, auto) | |
| 44123 | 1439 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1440 | |
| 53799 | 1441 | text {* Hence the following eccentric variant of the inverse function theorem.
 | 
| 1442 | This has no continuity assumptions, but we do need the inverse function. | |
| 1443 |   We could put @{text "f' \<circ> g = I"} but this happens to fit with the minimal linear
 | |
| 1444 | algebra theory I've set up so far. *} | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1445 | |
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36844diff
changeset | 1446 | (* 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 | 1447 | |
| 53781 | 1448 | lemma right_inverse_linear: | 
| 1449 | fixes f :: "'a::euclidean_space \<Rightarrow> 'a" | |
| 1450 | assumes lf: "linear f" | |
| 1451 | and gf: "f \<circ> g = id" | |
| 1452 | shows "linear g" | |
| 1453 | proof - | |
| 1454 | from gf have fi: "surj f" | |
| 1455 | by (auto simp add: surj_def o_def id_def) metis | |
| 1456 | from linear_surjective_isomorphism[OF lf fi] | |
| 1457 | obtain h:: "'a \<Rightarrow> 'a" where h: "linear h" "\<forall>x. h (f x) = x" "\<forall>x. f (h x) = x" | |
| 1458 | by blast | |
| 1459 | have "h = g" | |
| 1460 | apply (rule ext) | |
| 1461 | using gf h(2,3) | |
| 1462 | apply (simp add: o_def id_def fun_eq_iff) | |
| 1463 | apply metis | |
| 1464 | done | |
| 1465 | with h(1) show ?thesis by blast | |
| 1466 | qed | |
| 1467 | ||
| 44123 | 1468 | lemma has_derivative_inverse_strong: | 
| 53781 | 1469 | fixes f :: "'n::ordered_euclidean_space \<Rightarrow> 'n" | 
| 1470 | assumes "open s" | |
| 1471 | and "x \<in> s" | |
| 1472 | and "continuous_on s f" | |
| 1473 | and "\<forall>x\<in>s. g (f x) = x" | |
| 1474 | and "(f has_derivative f') (at x)" | |
| 1475 | and "f' \<circ> g' = id" | |
| 44123 | 1476 | shows "(g has_derivative g') (at (f x))" | 
| 53781 | 1477 | proof - | 
| 1478 | have linf: "bounded_linear f'" | |
| 44123 | 1479 | using assms(5) unfolding has_derivative_def by auto | 
| 53781 | 1480 | then have ling: "bounded_linear g'" | 
| 1481 | unfolding linear_conv_bounded_linear[symmetric] | |
| 1482 | apply - | |
| 1483 | apply (rule right_inverse_linear) | |
| 1484 | using assms(6) | |
| 1485 | apply auto | |
| 1486 | done | |
| 1487 | moreover have "g' \<circ> f' = id" | |
| 1488 | using assms(6) linf ling | |
| 1489 | unfolding linear_conv_bounded_linear[symmetric] | |
| 1490 | using linear_inverse_left | |
| 1491 | by auto | |
| 1492 | moreover have *:"\<forall>t\<subseteq>s. x \<in> interior t \<longrightarrow> f x \<in> interior (f ` t)" | |
| 1493 | apply rule | |
| 1494 | apply rule | |
| 1495 | apply rule | |
| 1496 | apply (rule sussmann_open_mapping) | |
| 1497 | apply (rule assms ling)+ | |
| 1498 | apply auto | |
| 1499 | done | |
| 1500 | have "continuous (at (f x)) g" | |
| 1501 | unfolding continuous_at Lim_at | |
| 1502 | proof (rule, rule) | |
| 1503 | fix e :: real | |
| 1504 | assume "e > 0" | |
| 1505 | then have "f x \<in> interior (f ` (ball x e \<inter> s))" | |
| 1506 | using *[rule_format,of "ball x e \<inter> s"] `x \<in> s` | |
| 1507 | by (auto simp add: interior_open[OF open_ball] interior_open[OF 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 | 1508 | then guess d unfolding mem_interior .. note d=this | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1509 | 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 | 1510 | apply (rule_tac x=d in exI) | 
| 1511 | apply rule | |
| 1512 | apply (rule d[THEN conjunct1]) | |
| 1513 | apply rule | |
| 1514 | apply rule | |
| 1515 | proof - | |
| 1516 | case goal1 | |
| 1517 | then have "g y \<in> g ` f ` (ball x e \<inter> s)" | |
| 44123 | 1518 | using d[THEN conjunct2,unfolded subset_eq,THEN bspec[where x=y]] | 
| 53781 | 1519 | by (auto simp add: dist_commute) | 
| 1520 | then have "g y \<in> ball x e \<inter> s" | |
| 1521 | using assms(4) by auto | |
| 1522 | then show "dist (g y) (g (f x)) < e" | |
| 1523 | using assms(4)[rule_format,OF `x \<in> s`] | |
| 44123 | 1524 | by (auto simp add: dist_commute) | 
| 1525 | qed | |
| 1526 | qed | |
| 1527 | moreover have "f x \<in> interior (f ` s)" | |
| 53781 | 1528 | apply (rule sussmann_open_mapping) | 
| 1529 | apply (rule assms ling)+ | |
| 1530 | using interior_open[OF assms(1)] and `x \<in> s` | |
| 1531 | apply auto | |
| 1532 | done | |
| 44123 | 1533 | moreover have "\<And>y. y \<in> interior (f ` s) \<Longrightarrow> f (g y) = y" | 
| 53781 | 1534 | proof - | 
| 1535 | case goal1 | |
| 1536 | then have "y \<in> f ` s" | |
| 1537 | using interior_subset by auto | |
| 44123 | 1538 | then guess z unfolding image_iff .. | 
| 53781 | 1539 | then show ?case | 
| 1540 | using assms(4) by auto | |
| 44123 | 1541 | qed | 
| 1542 | ultimately show ?thesis | |
| 53781 | 1543 | apply - | 
| 1544 | apply (rule has_derivative_inverse_basic_x[OF assms(5)]) | |
| 1545 | using assms | |
| 1546 | apply auto | |
| 1547 | done | |
| 44123 | 1548 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1549 | |
| 44124 
4c2a61a897d8
Derivative.thy: more sensible subsection headings
 huffman parents: 
44123diff
changeset | 1550 | text {* A rewrite based on the other domain. *}
 | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1551 | |
| 44123 | 1552 | lemma has_derivative_inverse_strong_x: | 
| 53781 | 1553 | fixes f :: "'a::ordered_euclidean_space \<Rightarrow> 'a" | 
| 1554 | assumes "open s" | |
| 1555 | and "g y \<in> s" | |
| 1556 | and "continuous_on s f" | |
| 1557 | and "\<forall>x\<in>s. g (f x) = x" | |
| 1558 | and "(f has_derivative f') (at (g y))" | |
| 1559 | and "f' \<circ> g' = id" | |
| 1560 | 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 | 1561 | shows "(g has_derivative g') (at y)" | 
| 53781 | 1562 | using has_derivative_inverse_strong[OF assms(1-6)] | 
| 1563 | unfolding assms(7) | |
| 1564 | 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 | 1565 | |
| 44124 
4c2a61a897d8
Derivative.thy: more sensible subsection headings
 huffman parents: 
44123diff
changeset | 1566 | text {* On a region. *}
 | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1567 | |
| 44123 | 1568 | lemma has_derivative_inverse_on: | 
| 53781 | 1569 | fixes f :: "'n::ordered_euclidean_space \<Rightarrow> 'n" | 
| 1570 | assumes "open s" | |
| 1571 | and "\<forall>x\<in>s. (f has_derivative f'(x)) (at x)" | |
| 1572 | and "\<forall>x\<in>s. g (f x) = x" | |
| 1573 | and "f' x \<circ> g' x = id" | |
| 1574 | 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 | 1575 | shows "(g has_derivative g'(x)) (at (f x))" | 
| 53781 | 1576 | apply (rule has_derivative_inverse_strong[where g'="g' x" and f=f]) | 
| 1577 | 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 | 1578 | unfolding continuous_on_eq_continuous_at[OF assms(1)] | 
| 53781 | 1579 | apply rule | 
| 1580 | apply (rule differentiable_imp_continuous_within) | |
| 1581 | unfolding differentiable_def | |
| 1582 | using assms | |
| 1583 | apply auto | |
| 1584 | done | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1585 | |
| 44123 | 1586 | text {* Invertible derivative continous at a point implies local
 | 
| 1587 | injectivity. It's only for this we need continuity of the derivative, | |
| 1588 | except of course if we want the fact that the inverse derivative is | |
| 1589 | also continuous. So if we know for some other reason that the inverse | |
| 1590 | function exists, it's OK. *} | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1591 | |
| 53781 | 1592 | lemma bounded_linear_sub: "bounded_linear f \<Longrightarrow> bounded_linear g \<Longrightarrow> bounded_linear (\<lambda>x. f x - g x)" | 
| 44123 | 1593 | using bounded_linear_add[of f "\<lambda>x. - g x"] bounded_linear_minus[of g] | 
| 1594 | by (auto simp add: algebra_simps) | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1595 | |
| 44123 | 1596 | lemma has_derivative_locally_injective: | 
| 53781 | 1597 | fixes f :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space" | 
| 1598 | assumes "a \<in> s" | |
| 1599 | and "open s" | |
| 1600 | and "bounded_linear g'" | |
| 1601 | and "g' \<circ> f' a = id" | |
| 1602 | and "\<forall>x\<in>s. (f has_derivative f' x) (at x)" | |
| 1603 | and "\<forall>e>0. \<exists>d>0. \<forall>x. dist a x < d \<longrightarrow> onorm (\<lambda>v. f' x v - f' a v) < e" | |
| 1604 | obtains t where "a \<in> t" "open t" "\<forall>x\<in>t. \<forall>x'\<in>t. f x' = f x \<longrightarrow> x' = x" | |
| 1605 | proof - | |
| 1606 | interpret bounded_linear g' | |
| 1607 | 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 | 1608 | note f'g' = assms(4)[unfolded id_def o_def,THEN cong] | 
| 53781 | 1609 | have "g' (f' a (\<Sum>Basis)) = (\<Sum>Basis)" "(\<Sum>Basis) \<noteq> (0::'n)" | 
| 1610 | defer | |
| 1611 | apply (subst euclidean_eq_iff) | |
| 1612 | using f'g' | |
| 1613 | apply auto | |
| 1614 | done | |
| 1615 | then have *: "0 < onorm g'" | |
| 1616 | unfolding onorm_pos_lt[OF assms(3)[unfolded linear_linear]] | |
| 1617 | by fastforce | |
| 1618 | def k \<equiv> "1 / onorm g' / 2" | |
| 1619 | have *: "k > 0" | |
| 1620 | unfolding k_def using * 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 | 1621 | guess d1 using assms(6)[rule_format,OF *] .. note d1=this | 
| 53781 | 1622 | from `open s` obtain d2 where "d2 > 0" "ball a d2 \<subseteq> s" | 
| 1623 | using `a\<in>s` .. | |
| 1624 | obtain d2 where "d2 > 0" "ball a d2 \<subseteq> s" | |
| 1625 | using assms(2,1) .. | |
| 44123 | 1626 | guess d2 using assms(2)[unfolded open_contains_ball,rule_format,OF `a\<in>s`] .. | 
| 1627 | note d2=this | |
| 1628 | guess d using real_lbound_gt_zero[OF d1[THEN conjunct1] d2[THEN conjunct1]] .. | |
| 1629 | note d = this | |
| 1630 | show ?thesis | |
| 1631 | proof | |
| 53781 | 1632 | show "a \<in> ball a d" | 
| 1633 | using d by auto | |
| 44123 | 1634 | show "\<forall>x\<in>ball a d. \<forall>x'\<in>ball a d. f x' = f x \<longrightarrow> x' = x" | 
| 1635 | proof (intro strip) | |
| 53781 | 1636 | fix x y | 
| 1637 | assume as: "x \<in> ball a d" "y \<in> ball a d" "f x = f y" | |
| 1638 | def ph \<equiv> "\<lambda>w. w - g' (f w - f x)" | |
| 44123 | 1639 | have ph':"ph = g' \<circ> (\<lambda>w. f' a w - (f w - f x))" | 
| 53781 | 1640 | unfolding ph_def o_def | 
| 1641 | unfolding diff | |
| 1642 | using f'g' | |
| 44123 | 1643 | by (auto simp add: algebra_simps) | 
| 53781 | 1644 | have "norm (ph x - ph y) \<le> (1 / 2) * norm (x - y)" | 
| 1645 | apply (rule differentiable_bound[OF convex_ball _ _ as(1-2), where f'="\<lambda>x v. v - g'(f' x v)"]) | |
| 1646 | apply (rule_tac[!] ballI) | |
| 1647 | proof - | |
| 1648 | fix u | |
| 1649 | assume u: "u \<in> ball a d" | |
| 1650 | then have "u \<in> s" | |
| 1651 | using d d2 by auto | |
| 1652 | have *: "(\<lambda>v. v - g' (f' u v)) = g' \<circ> (\<lambda>w. f' a w - f' u w)" | |
| 1653 | unfolding o_def and diff | |
| 1654 | using f'g' by auto | |
| 41958 | 1655 | 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 | 1656 | unfolding ph' * | 
| 53781 | 1657 | apply (simp add: comp_def) | 
| 1658 | apply (rule bounded_linear.FDERIV[OF assms(3)]) | |
| 1659 | apply (rule FDERIV_intros) | |
| 1660 | defer | |
| 1661 | apply (rule has_derivative_sub[where g'="\<lambda>x.0",unfolded diff_0_right]) | |
| 1662 | apply (rule has_derivative_at_within) | |
| 1663 | using assms(5) and `u \<in> s` `a \<in> s` | |
| 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 | 1664 | apply (auto intro!: FDERIV_intros bounded_linear.FDERIV[of _ "\<lambda>x. x"] derivative_linear) | 
| 
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 | 1665 | done | 
| 53781 | 1666 | have **: "bounded_linear (\<lambda>x. f' u x - f' a x)" "bounded_linear (\<lambda>x. f' a x - f' u x)" | 
| 1667 | apply (rule_tac[!] bounded_linear_sub) | |
| 1668 | apply (rule_tac[!] derivative_linear) | |
| 1669 | using assms(5) `u \<in> s` `a \<in> s` | |
| 1670 | apply auto | |
| 1671 | done | |
| 44123 | 1672 | have "onorm (\<lambda>v. v - g' (f' u v)) \<le> onorm g' * onorm (\<lambda>w. f' a w - f' u w)" | 
| 53781 | 1673 | unfolding * | 
| 1674 | apply (rule onorm_compose) | |
| 1675 | unfolding linear_conv_bounded_linear | |
| 1676 | apply (rule assms(3) **)+ | |
| 1677 | done | |
| 44123 | 1678 | also have "\<dots> \<le> onorm g' * k" | 
| 53781 | 1679 | apply (rule mult_left_mono) | 
| 44123 | 1680 | using d1[THEN conjunct2,rule_format,of u] | 
| 1681 | using onorm_neg[OF **(1)[unfolded linear_linear]] | |
| 1682 | using d and u and onorm_pos_le[OF assms(3)[unfolded linear_linear]] | |
| 53781 | 1683 | apply (auto simp add: algebra_simps) | 
| 1684 | done | |
| 1685 | also have "\<dots> \<le> 1 / 2" | |
| 1686 | unfolding k_def by auto | |
| 1687 | finally show "onorm (\<lambda>v. v - g' (f' u v)) \<le> 1 / 2" . | |
| 44123 | 1688 | qed | 
| 1689 | moreover have "norm (ph y - ph x) = norm (y - x)" | |
| 53781 | 1690 | apply (rule arg_cong[where f=norm]) | 
| 1691 | unfolding ph_def | |
| 1692 | using diff | |
| 1693 | unfolding as | |
| 1694 | apply auto | |
| 1695 | done | |
| 1696 | ultimately show "x = y" | |
| 1697 | unfolding norm_minus_commute by auto | |
| 44123 | 1698 | qed | 
| 1699 | qed auto | |
| 1700 | qed | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1701 | |
| 53781 | 1702 | |
| 1703 | subsection {* Uniformly convergent sequence of derivatives *}
 | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1704 | |
| 44123 | 1705 | lemma has_derivative_sequence_lipschitz_lemma: | 
| 53781 | 1706 | fixes f :: "nat \<Rightarrow> 'm::euclidean_space \<Rightarrow> 'n::euclidean_space" | 
| 44123 | 1707 | assumes "convex s" | 
| 53781 | 1708 | and "\<forall>n. \<forall>x\<in>s. ((f n) has_derivative (f' n x)) (at x within s)" | 
| 1709 | and "\<forall>n\<ge>N. \<forall>x\<in>s. \<forall>h. norm (f' n x h - g' x h) \<le> e * norm h" | |
| 1710 | 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)" | |
| 1711 | proof rule+ | |
| 1712 | fix m n x y | |
| 1713 | assume as: "N \<le> m" "N \<le> n" "x \<in> s" "y \<in> s" | |
| 1714 | show "norm ((f m x - f n x) - (f m y - f n y)) \<le> 2 * e * norm (x - y)" | |
| 1715 | apply (rule differentiable_bound[where f'="\<lambda>x h. f' m x h - f' n x h", OF assms(1) _ _ as(3-4)]) | |
| 1716 | apply (rule_tac[!] ballI) | |
| 1717 | proof - | |
| 1718 | fix x | |
| 1719 | assume "x \<in> s" | |
| 44123 | 1720 | 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)" | 
| 53781 | 1721 | by (rule FDERIV_intros assms(2)[rule_format] `x\<in>s`)+ | 
| 1722 |     {
 | |
| 1723 | fix h | |
| 44123 | 1724 | 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)" | 
| 1725 | using norm_triangle_ineq[of "f' m x h - g' x h" "- f' n x h + g' x h"] | |
| 53781 | 1726 | unfolding norm_minus_commute | 
| 1727 | by (auto simp add: algebra_simps) | |
| 1728 | also have "\<dots> \<le> e * norm h + e * norm h" | |
| 1729 | using assms(3)[rule_format,OF `N \<le> m` `x \<in> s`, of h] | |
| 1730 | using assms(3)[rule_format,OF `N \<le> n` `x \<in> s`, of h] | |
| 1731 | by (auto simp add: field_simps) | |
| 1732 | finally have "norm (f' m x h - f' n x h) \<le> 2 * e * norm h" | |
| 1733 | by auto | |
| 1734 | } | |
| 1735 | then show "onorm (\<lambda>h. f' m x h - f' n x h) \<le> 2 * e" | |
| 1736 | apply - | |
| 1737 | apply (rule onorm(2)) | |
| 1738 | apply (rule linear_compose_sub) | |
| 44123 | 1739 | unfolding linear_conv_bounded_linear | 
| 53781 | 1740 | using assms(2)[rule_format,OF `x \<in> s`, THEN derivative_linear] | 
| 1741 | apply auto | |
| 1742 | done | |
| 44123 | 1743 | qed | 
| 1744 | qed | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1745 | |
| 44123 | 1746 | lemma has_derivative_sequence_lipschitz: | 
| 53781 | 1747 | fixes f :: "nat \<Rightarrow> 'm::euclidean_space \<Rightarrow> 'n::euclidean_space" | 
| 44123 | 1748 | assumes "convex s" | 
| 53781 | 1749 | and "\<forall>n. \<forall>x\<in>s. ((f n) has_derivative (f' n x)) (at x within s)" | 
| 1750 | 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" | |
| 1751 | and "e > 0" | |
| 1752 | shows "\<forall>e>0. \<exists>N. \<forall>m\<ge>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>y\<in>s. | |
| 1753 | norm ((f m x - f n x) - (f m y - f n y)) \<le> e * norm (x - y)" | |
| 1754 | proof (rule, rule) | |
| 1755 | case goal1 have *: "2 * (1/2* e) = e" "1/2 * e >0" | |
| 1756 | using `e > 0` 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 | 1757 | guess N using assms(3)[rule_format,OF *(2)] .. | 
| 53781 | 1758 | then show ?case | 
| 1759 | apply (rule_tac x=N in exI) | |
| 1760 | apply (rule has_derivative_sequence_lipschitz_lemma[where e="1/2 *e", unfolded *]) | |
| 1761 | using assms | |
| 1762 | apply auto | |
| 1763 | done | |
| 44123 | 1764 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1765 | |
| 44123 | 1766 | lemma has_derivative_sequence: | 
| 1767 | fixes f::"nat\<Rightarrow> 'm::euclidean_space \<Rightarrow> 'n::euclidean_space" | |
| 1768 | assumes "convex s" | |
| 53781 | 1769 | and "\<forall>n. \<forall>x\<in>s. ((f n) has_derivative (f' n x)) (at x within s)" | 
| 1770 | 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" | |
| 1771 | and "x0 \<in> s" | |
| 1772 | and "((\<lambda>n. f n x0) ---> l) sequentially" | |
| 1773 | shows "\<exists>g. \<forall>x\<in>s. ((\<lambda>n. f n x) ---> g x) sequentially \<and> (g has_derivative g'(x)) (at x within s)" | |
| 1774 | proof - | |
| 1775 | have lem1: "\<forall>e>0. \<exists>N. \<forall>m\<ge>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>y\<in>s. | |
| 1776 | norm ((f m x - f n x) - (f m y - f n y)) \<le> e * norm (x - y)" | |
| 1777 | apply (rule has_derivative_sequence_lipschitz[where e="42::nat"]) | |
| 1778 | apply (rule assms)+ | |
| 1779 | apply auto | |
| 1780 | done | |
| 44123 | 1781 | have "\<exists>g. \<forall>x\<in>s. ((\<lambda>n. f n x) ---> g x) sequentially" | 
| 53781 | 1782 | apply (rule bchoice) | 
| 1783 | unfolding convergent_eq_cauchy | |
| 44123 | 1784 | proof | 
| 53781 | 1785 | fix x | 
| 1786 | assume "x \<in> s" | |
| 1787 | show "Cauchy (\<lambda>n. f n x)" | |
| 1788 | proof (cases "x = x0") | |
| 1789 | case True | |
| 1790 | then show ?thesis | |
| 1791 | using LIMSEQ_imp_Cauchy[OF assms(5)] by auto | |
| 44123 | 1792 | next | 
| 53781 | 1793 | case False | 
| 1794 | show ?thesis | |
| 1795 | unfolding Cauchy_def | |
| 1796 | proof (rule, rule) | |
| 1797 | fix e :: real | |
| 1798 | assume "e > 0" | |
| 1799 | then have *: "e / 2 > 0" "e / 2 / norm (x - x0) > 0" | |
| 44123 | 1800 | using False by (auto intro!: divide_pos_pos) | 
| 50939 
ae7cd20ed118
replace convergent_imp_cauchy by LIMSEQ_imp_Cauchy
 hoelzl parents: 
50526diff
changeset | 1801 | guess M using LIMSEQ_imp_Cauchy[OF assms(5), unfolded Cauchy_def, rule_format,OF *(1)] .. note M=this | 
| 41958 | 1802 | guess N using lem1[rule_format,OF *(2)] .. note N = this | 
| 44123 | 1803 | show "\<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. dist (f m x) (f n x) < e" | 
| 53781 | 1804 | apply (rule_tac x="max M N" in exI) | 
| 1805 | proof rule+ | |
| 1806 | fix m n | |
| 1807 | assume as: "max M N \<le>m" "max M N\<le>n" | |
| 1808 | have "dist (f m x) (f n x) \<le> | |
| 1809 | norm (f m x0 - f n x0) + norm (f m x - f n x - (f m x0 - f n x0))" | |
| 1810 | unfolding dist_norm | |
| 1811 | by (rule norm_triangle_sub) | |
| 44123 | 1812 | also have "\<dots> \<le> norm (f m x0 - f n x0) + e / 2" | 
| 1813 | using N[rule_format,OF _ _ `x\<in>s` `x0\<in>s`, of m n] and as and False | |
| 1814 | by auto | |
| 1815 | also have "\<dots> < e / 2 + e / 2" | |
| 53781 | 1816 | apply (rule add_strict_right_mono) | 
| 1817 | using as and M[rule_format] | |
| 1818 | unfolding dist_norm | |
| 1819 | apply auto | |
| 1820 | done | |
| 1821 | finally show "dist (f m x) (f n x) < e" | |
| 1822 | by auto | |
| 44123 | 1823 | qed | 
| 1824 | qed | |
| 1825 | qed | |
| 1826 | qed | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1827 | then guess g .. note g = this | 
| 53781 | 1828 | 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)" | 
| 1829 | proof (rule, rule) | |
| 1830 | fix e :: real | |
| 1831 | assume *: "e > 0" | |
| 44123 | 1832 | guess N using lem1[rule_format,OF *] .. note N=this | 
| 1833 | 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 | 1834 | apply (rule_tac x=N in exI) | 
| 1835 | proof rule+ | |
| 1836 | fix n x y | |
| 1837 | assume as: "N \<le> n" "x \<in> s" "y \<in> s" | |
| 44123 | 1838 | have "eventually (\<lambda>xa. norm (f n x - f n y - (f xa x - f xa y)) \<le> e * norm (x - y)) sequentially" | 
| 1839 | unfolding eventually_sequentially | |
| 53781 | 1840 | apply (rule_tac x=N in exI) | 
| 1841 | apply rule | |
| 1842 | apply rule | |
| 1843 | proof - | |
| 1844 | fix m | |
| 1845 | assume "N \<le> m" | |
| 1846 | then show "norm (f n x - f n y - (f m x - f m y)) \<le> e * norm (x - y)" | |
| 44123 | 1847 | using N[rule_format, of n m x y] and as | 
| 1848 | by (auto simp add: algebra_simps) | |
| 1849 | qed | |
| 53781 | 1850 | then show "norm (f n x - f n y - (g x - g y)) \<le> e * norm (x - y)" | 
| 1851 | apply - | |
| 1852 | apply (rule Lim_norm_ubound[OF trivial_limit_sequentially, where f="\<lambda>m. (f n x - f n y) - (f m x - f m y)"]) | |
| 1853 | apply (rule tendsto_intros g[rule_format] as)+ | |
| 1854 | apply assumption | |
| 1855 | done | |
| 44123 | 1856 | qed | 
| 1857 | qed | |
| 53781 | 1858 | show ?thesis | 
| 1859 | unfolding has_derivative_within_alt | |
| 1860 | apply (rule_tac x=g in exI) | |
| 1861 | apply rule | |
| 1862 | apply rule | |
| 1863 | apply (rule g[rule_format]) | |
| 1864 | apply assumption | |
| 1865 | proof | |
| 1866 | fix x | |
| 1867 | assume "x \<in> s" | |
| 1868 | have lem3: "\<forall>u. ((\<lambda>n. f' n x u) ---> g' x u) sequentially" | |
| 44907 
93943da0a010
remove redundant lemma Lim_sequentially in favor of lemma LIMSEQ_def
 huffman parents: 
44890diff
changeset | 1869 | unfolding LIMSEQ_def | 
| 53781 | 1870 | proof (rule, rule, rule) | 
| 1871 | fix u | |
| 1872 | fix e :: real | |
| 1873 | assume "e > 0" | |
| 44123 | 1874 | show "\<exists>N. \<forall>n\<ge>N. dist (f' n x u) (g' x u) < e" | 
| 53781 | 1875 | proof (cases "u = 0") | 
| 1876 | case True | |
| 1877 | guess N using assms(3)[rule_format,OF `e>0`] .. note N=this | |
| 1878 | show ?thesis | |
| 1879 | apply (rule_tac x=N in exI) | |
| 1880 | unfolding True | |
| 1881 | using N[rule_format,OF _ `x\<in>s`,of _ 0] and `e>0` | |
| 1882 | apply auto | |
| 1883 | done | |
| 44123 | 1884 | next | 
| 53781 | 1885 | case False | 
| 1886 | then have *: "e / 2 / norm u > 0" | |
| 1887 | using `e > 0` | |
| 1888 | by (auto intro!: divide_pos_pos) | |
| 41958 | 1889 | guess N using assms(3)[rule_format,OF *] .. note N=this | 
| 53781 | 1890 | show ?thesis | 
| 1891 | apply (rule_tac x=N in exI) | |
| 1892 | apply rule | |
| 1893 | apply rule | |
| 1894 | proof - | |
| 1895 | case goal1 | |
| 1896 | show ?case | |
| 1897 | unfolding dist_norm | |
| 44123 | 1898 | using N[rule_format,OF goal1 `x\<in>s`, of u] False `e>0` | 
| 53781 | 1899 | by (auto simp add: field_simps) | 
| 44123 | 1900 | qed | 
| 1901 | qed | |
| 1902 | qed | |
| 1903 | show "bounded_linear (g' x)" | |
| 53600 
8fda7ad57466
make 'linear' into a sublocale of 'bounded_linear';
 huffman parents: 
53374diff
changeset | 1904 | unfolding linear_linear linear_iff | 
| 53781 | 1905 | apply rule | 
| 1906 | apply rule | |
| 1907 | apply rule | |
| 1908 | defer | |
| 1909 | apply rule | |
| 1910 | apply rule | |
| 1911 | proof - | |
| 1912 | fix x' y z :: 'm | |
| 1913 | fix c :: real | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1914 | note lin = assms(2)[rule_format,OF `x\<in>s`,THEN derivative_linear] | 
| 44123 | 1915 | show "g' x (c *\<^sub>R x') = c *\<^sub>R g' x x'" | 
| 53781 | 1916 | apply (rule tendsto_unique[OF trivial_limit_sequentially]) | 
| 1917 | apply (rule lem3[rule_format]) | |
| 53600 
8fda7ad57466
make 'linear' into a sublocale of 'bounded_linear';
 huffman parents: 
53374diff
changeset | 1918 | unfolding lin[THEN bounded_linear_imp_linear, THEN linear_cmul] | 
| 53781 | 1919 | apply (intro tendsto_intros) | 
| 1920 | apply (rule lem3[rule_format]) | |
| 1921 | done | |
| 44123 | 1922 | show "g' x (y + z) = g' x y + g' x z" | 
| 53781 | 1923 | apply (rule tendsto_unique[OF trivial_limit_sequentially]) | 
| 1924 | apply (rule lem3[rule_format]) | |
| 53600 
8fda7ad57466
make 'linear' into a sublocale of 'bounded_linear';
 huffman parents: 
53374diff
changeset | 1925 | unfolding lin[THEN bounded_linear_imp_linear, THEN linear_add] | 
| 53781 | 1926 | apply (rule tendsto_add) | 
| 1927 | apply (rule lem3[rule_format])+ | |
| 1928 | done | |
| 44123 | 1929 | qed | 
| 1930 | show "\<forall>e>0. \<exists>d>0. \<forall>y\<in>s. norm (y - x) < d \<longrightarrow> norm (g y - g x - g' x (y - x)) \<le> e * norm (y - x)" | |
| 53781 | 1931 | proof (rule, rule) | 
| 1932 | case goal1 | |
| 1933 | have *: "e / 3 > 0" | |
| 1934 | using goal1 by auto | |
| 44123 | 1935 | guess N1 using assms(3)[rule_format,OF *] .. note N1=this | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1936 | guess N2 using lem2[rule_format,OF *] .. note N2=this | 
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1937 | guess d1 using assms(2)[unfolded has_derivative_within_alt, rule_format,OF `x\<in>s`, of "max N1 N2",THEN conjunct2,rule_format,OF *] .. note d1=this | 
| 53781 | 1938 | show ?case | 
| 1939 | apply (rule_tac x=d1 in exI) | |
| 1940 | apply rule | |
| 1941 | apply (rule d1[THEN conjunct1]) | |
| 1942 | apply rule | |
| 1943 | apply rule | |
| 1944 | proof - | |
| 1945 | fix y | |
| 1946 | assume as: "y \<in> s" "norm (y - x) < d1" | |
| 1947 | let ?N = "max N1 N2" | |
| 44123 | 1948 | have "norm (g y - g x - (f ?N y - f ?N x)) \<le> e /3 * norm (y - x)" | 
| 53781 | 1949 | apply (subst norm_minus_cancel[symmetric]) | 
| 1950 | using N2[rule_format, OF _ `y \<in> s` `x \<in> s`, of ?N] | |
| 1951 | apply auto | |
| 1952 | done | |
| 44123 | 1953 | moreover | 
| 1954 | have "norm (f ?N y - f ?N x - f' ?N x (y - x)) \<le> e / 3 * norm (y - x)" | |
| 53781 | 1955 | using d1 and as | 
| 1956 | by auto | |
| 44123 | 1957 | ultimately | 
| 53781 | 1958 | have "norm (g y - g x - f' ?N x (y - x)) \<le> 2 * e / 3 * norm (y - x)" | 
| 44123 | 1959 | 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 | 1960 | by (auto simp add: algebra_simps) | 
| 44123 | 1961 | moreover | 
| 1962 | have " norm (f' ?N x (y - x) - g' x (y - x)) \<le> e / 3 * norm (y - x)" | |
| 53781 | 1963 | using N1 `x \<in> s` by auto | 
| 41958 | 1964 | ultimately show "norm (g y - g x - g' x (y - x)) \<le> e * norm (y - x)" | 
| 44123 | 1965 | 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 | 1966 | by (auto simp add: algebra_simps) | 
| 44123 | 1967 | qed | 
| 1968 | qed | |
| 1969 | qed | |
| 1970 | qed | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1971 | |
| 44124 
4c2a61a897d8
Derivative.thy: more sensible subsection headings
 huffman parents: 
44123diff
changeset | 1972 | text {* Can choose to line up antiderivatives if we want. *}
 | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 1973 | |
| 44123 | 1974 | lemma has_antiderivative_sequence: | 
| 53781 | 1975 | fixes f :: "nat \<Rightarrow> 'm::euclidean_space \<Rightarrow> 'n::euclidean_space" | 
| 44123 | 1976 | assumes "convex s" | 
| 53781 | 1977 | and "\<forall>n. \<forall>x\<in>s. ((f n) has_derivative (f' n x)) (at x within s)" | 
| 1978 | 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" | |
| 1979 | shows "\<exists>g. \<forall>x\<in>s. (g has_derivative g' x) (at x within s)" | |
| 1980 | proof (cases "s = {}")
 | |
| 1981 | case False | |
| 1982 | then obtain a where "a \<in> s" | |
| 1983 | by auto | |
| 1984 | 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" | |
| 1985 | by auto | |
| 44123 | 1986 | show ?thesis | 
| 53781 | 1987 | apply (rule *) | 
| 1988 | apply (rule has_derivative_sequence[OF assms(1) _ assms(3), of "\<lambda>n x. f n x + (f 0 a - f n a)"]) | |
| 1989 | apply rule | |
| 1990 | apply rule | |
| 1991 | apply (rule has_derivative_add_const, rule assms(2)[rule_format]) | |
| 1992 | apply assumption | |
| 1993 | apply (rule `a \<in> s`) | |
| 1994 | apply auto | |
| 1995 | done | |
| 44123 | 1996 | 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 | 1997 | |
| 44123 | 1998 | lemma has_antiderivative_limit: | 
| 53781 | 1999 | fixes g' :: "'m::euclidean_space \<Rightarrow> 'm \<Rightarrow> 'n::euclidean_space" | 
| 44123 | 2000 | assumes "convex s" | 
| 53781 | 2001 | and "\<forall>e>0. \<exists>f f'. \<forall>x\<in>s. | 
| 2002 | (f has_derivative (f' x)) (at x within s) \<and> (\<forall>h. norm (f' x h - g' x h) \<le> e * norm h)" | |
| 2003 | shows "\<exists>g. \<forall>x\<in>s. (g has_derivative g' x) (at x within s)" | |
| 2004 | proof - | |
| 2005 | have *: "\<forall>n. \<exists>f f'. \<forall>x\<in>s. | |
| 2006 | (f has_derivative (f' x)) (at x within s) \<and> | |
| 2007 | (\<forall>h. norm(f' x h - g' x h) \<le> inverse (real (Suc n)) * norm h)" | |
| 2008 | apply rule | |
| 2009 | using assms(2) | |
| 2010 | apply (erule_tac x="inverse (real (Suc n))" in allE) | |
| 2011 | apply auto | |
| 2012 | done | |
| 44123 | 2013 | guess f using *[THEN choice] .. note * = this | 
| 53781 | 2014 | guess f' using *[THEN choice] .. note f = this | 
| 2015 | show ?thesis | |
| 2016 | apply (rule has_antiderivative_sequence[OF assms(1), of f f']) | |
| 2017 | defer | |
| 2018 | apply rule | |
| 2019 | apply rule | |
| 2020 | proof - | |
| 2021 | fix e :: real | |
| 2022 | assume "e > 0" | |
| 2023 | guess N using reals_Archimedean[OF `e>0`] .. note N=this | |
| 44123 | 2024 | 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 | 2025 | apply (rule_tac x=N in exI) | 
| 2026 | proof rule+ | |
| 44123 | 2027 | case goal1 | 
| 53781 | 2028 | have *: "inverse (real (Suc n)) \<le> e" | 
| 2029 | apply (rule order_trans[OF _ N[THEN less_imp_le]]) | |
| 2030 | using goal1(1) | |
| 2031 | apply (auto simp add: field_simps) | |
| 2032 | done | |
| 44123 | 2033 | show ?case | 
| 53781 | 2034 | using f[rule_format,THEN conjunct2,OF goal1(2), of n, THEN spec[where x=h]] | 
| 2035 | apply (rule order_trans) | |
| 2036 | using N * | |
| 2037 | apply (cases "h = 0") | |
| 2038 | apply auto | |
| 2039 | done | |
| 44123 | 2040 | qed | 
| 53781 | 2041 | qed (insert f, auto) | 
| 44123 | 2042 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2043 | |
| 53781 | 2044 | |
| 2045 | subsection {* Differentiation of a series *}
 | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2046 | |
| 53781 | 2047 | definition sums_seq :: "(nat \<Rightarrow> 'a::real_normed_vector) \<Rightarrow> 'a \<Rightarrow> nat set \<Rightarrow> bool" | 
| 2048 | (infixl "sums'_seq" 12) | |
| 2049 |   where "(f sums_seq l) s \<longleftrightarrow> ((\<lambda>n. setsum f (s \<inter> {0..n})) ---> l) sequentially"
 | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2050 | |
| 44123 | 2051 | lemma has_derivative_series: | 
| 53781 | 2052 | fixes f :: "nat \<Rightarrow> 'm::euclidean_space \<Rightarrow> 'n::euclidean_space" | 
| 44123 | 2053 | assumes "convex s" | 
| 53781 | 2054 | and "\<forall>n. \<forall>x\<in>s. ((f n) has_derivative (f' n x)) (at x within s)" | 
| 2055 |     and "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>h. norm (setsum (\<lambda>i. f' i x h) (k \<inter> {0..n}) - g' x h) \<le> e * norm h"
 | |
| 2056 | and "x \<in> s" | |
| 2057 | and "((\<lambda>n. f n x) sums_seq l) k" | |
| 2058 | shows "\<exists>g. \<forall>x\<in>s. ((\<lambda>n. f n x) sums_seq (g x)) k \<and> (g has_derivative g' x) (at x within s)" | |
| 44123 | 2059 | unfolding sums_seq_def | 
| 53781 | 2060 | apply (rule has_derivative_sequence[OF assms(1) _ assms(3)]) | 
| 2061 | apply rule | |
| 2062 | apply rule | |
| 2063 | apply (rule has_derivative_setsum) | |
| 2064 | apply (rule assms(2)[rule_format]) | |
| 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 | 2065 | apply assumption | 
| 53781 | 2066 | using assms(4-5) | 
| 2067 | unfolding sums_seq_def | |
| 2068 | apply auto | |
| 2069 | done | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2070 | |
| 53781 | 2071 | |
| 2072 | text {* Considering derivative @{typ "real \<Rightarrow> 'b\<Colon>real_normed_vector"} as a vector. *}
 | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2073 | |
| 53781 | 2074 | definition has_vector_derivative :: "(real \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> 'b \<Rightarrow> real filter \<Rightarrow> bool" | 
| 2075 | (infixl "has'_vector'_derivative" 12) | |
| 2076 | where "(f has_vector_derivative f') net \<longleftrightarrow> (f has_derivative (\<lambda>x. x *\<^sub>R 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 | 2077 | |
| 53781 | 2078 | definition "vector_derivative f net = (SOME f'. (f has_vector_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 | 2079 | |
| 44123 | 2080 | lemma vector_derivative_works: | 
| 53781 | 2081 | fixes f :: "real \<Rightarrow> 'a::real_normed_vector" | 
| 2082 | shows "f differentiable net \<longleftrightarrow> (f has_vector_derivative (vector_derivative f net)) net" | |
| 2083 | (is "?l = ?r") | |
| 44123 | 2084 | proof | 
| 53781 | 2085 | assume ?l | 
| 2086 | guess f' using `?l`[unfolded differentiable_def] .. note f' = this | |
| 2087 | then interpret bounded_linear f' | |
| 2088 | by auto | |
| 2089 | show ?r | |
| 2090 | unfolding vector_derivative_def has_vector_derivative_def | |
| 2091 | apply - | |
| 2092 | apply (rule someI_ex,rule_tac x="f' 1" in exI) | |
| 2093 | using f' | |
| 2094 | unfolding scaleR[symmetric] | |
| 2095 | apply auto | |
| 2096 | done | |
| 44123 | 2097 | next | 
| 53781 | 2098 | assume ?r | 
| 2099 | then show ?l | |
| 44123 | 2100 | unfolding vector_derivative_def has_vector_derivative_def differentiable_def | 
| 2101 | by auto | |
| 2102 | qed | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2103 | |
| 50418 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 hoelzl parents: 
46898diff
changeset | 2104 | lemma has_vector_derivative_withinI_DERIV: | 
| 53781 | 2105 | assumes f: "DERIV f x :> y" | 
| 2106 | shows "(f has_vector_derivative y) (at x within s)" | |
| 50418 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 hoelzl parents: 
46898diff
changeset | 2107 | unfolding has_vector_derivative_def real_scaleR_def | 
| 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 hoelzl parents: 
46898diff
changeset | 2108 | apply (rule has_derivative_at_within) | 
| 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 hoelzl parents: 
46898diff
changeset | 2109 | using DERIV_conv_has_derivative[THEN iffD1, OF f] | 
| 53781 | 2110 | apply (subst mult_commute) | 
| 2111 | apply assumption | |
| 2112 | done | |
| 50418 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 hoelzl parents: 
46898diff
changeset | 2113 | |
| 37730 | 2114 | lemma vector_derivative_unique_at: | 
| 2115 | assumes "(f has_vector_derivative f') (at x)" | |
| 53781 | 2116 | and "(f has_vector_derivative f'') (at x)" | 
| 37730 | 2117 | shows "f' = f''" | 
| 53781 | 2118 | proof - | 
| 37730 | 2119 | have "(\<lambda>x. x *\<^sub>R f') = (\<lambda>x. x *\<^sub>R f'')" | 
| 2120 | using assms [unfolded has_vector_derivative_def] | |
| 2121 | by (rule frechet_derivative_unique_at) | |
| 53781 | 2122 | then show ?thesis | 
| 2123 | unfolding fun_eq_iff by auto | |
| 37730 | 2124 | qed | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2125 | |
| 44123 | 2126 | lemma vector_derivative_unique_within_closed_interval: | 
| 53781 | 2127 | assumes "a < b" | 
| 2128 |     and "x \<in> {a..b}"
 | |
| 44123 | 2129 |   assumes "(f has_vector_derivative f') (at x within {a..b})"
 | 
| 2130 |   assumes "(f has_vector_derivative f'') (at x within {a..b})"
 | |
| 2131 | shows "f' = f''" | |
| 53781 | 2132 | proof - | 
| 2133 | have *: "(\<lambda>x. x *\<^sub>R f') = (\<lambda>x. x *\<^sub>R f'')" | |
| 2134 | apply (rule frechet_derivative_unique_within_closed_interval[of "a" "b"]) | |
| 2135 | using assms(3-)[unfolded has_vector_derivative_def] | |
| 2136 | using assms(1-2) | |
| 2137 | apply auto | |
| 2138 | done | |
| 44123 | 2139 | show ?thesis | 
| 53781 | 2140 | proof (rule ccontr) | 
| 53374 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
53077diff
changeset | 2141 | assume **: "f' \<noteq> f''" | 
| 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
53077diff
changeset | 2142 | with * have "(\<lambda>x. x *\<^sub>R f') 1 = (\<lambda>x. x *\<^sub>R f'') 1" | 
| 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
53077diff
changeset | 2143 | by (auto simp: fun_eq_iff) | 
| 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
53077diff
changeset | 2144 | with ** show False | 
| 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
53077diff
changeset | 2145 | unfolding o_def by auto | 
| 44123 | 2146 | qed | 
| 2147 | qed | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2148 | |
| 37730 | 2149 | lemma vector_derivative_at: | 
| 53781 | 2150 | "(f has_vector_derivative f') (at x) \<Longrightarrow> vector_derivative f (at x) = f'" | 
| 2151 | apply (rule vector_derivative_unique_at) | |
| 2152 | defer | |
| 2153 | apply assumption | |
| 2154 | unfolding vector_derivative_works[symmetric] differentiable_def | |
| 2155 | unfolding has_vector_derivative_def | |
| 2156 | apply auto | |
| 2157 | done | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2158 | |
| 44123 | 2159 | lemma vector_derivative_within_closed_interval: | 
| 53781 | 2160 | assumes "a < b" | 
| 2161 |     and "x \<in> {a..b}"
 | |
| 44123 | 2162 |   assumes "(f has_vector_derivative f') (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 | 2163 |   shows "vector_derivative f (at x within {a..b}) = f'"
 | 
| 53781 | 2164 | apply (rule vector_derivative_unique_within_closed_interval) | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2165 | using vector_derivative_works[unfolded differentiable_def] | 
| 53781 | 2166 | using assms | 
| 2167 | apply (auto simp add:has_vector_derivative_def) | |
| 2168 | done | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2169 | |
| 53781 | 2170 | lemma has_vector_derivative_within_subset: | 
| 2171 | "(f has_vector_derivative f') (at x within s) \<Longrightarrow> t \<subseteq> s \<Longrightarrow> | |
| 2172 | (f has_vector_derivative f') (at x within t)" | |
| 2173 | unfolding has_vector_derivative_def | |
| 2174 | apply (rule has_derivative_within_subset) | |
| 2175 | apply auto | |
| 2176 | done | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2177 | |
| 53781 | 2178 | lemma has_vector_derivative_const: "((\<lambda>x. c) has_vector_derivative 0) net" | 
| 2179 | unfolding has_vector_derivative_def | |
| 2180 | using has_derivative_const | |
| 2181 | 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 | 2182 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2183 | lemma has_vector_derivative_id: "((\<lambda>x::real. x) has_vector_derivative 1) net" | 
| 53781 | 2184 | unfolding has_vector_derivative_def | 
| 2185 | using has_derivative_id | |
| 2186 | 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 | 2187 | |
| 44123 | 2188 | lemma has_vector_derivative_cmul: | 
| 53781 | 2189 | "(f has_vector_derivative f') net \<Longrightarrow> | 
| 2190 | ((\<lambda>x. c *\<^sub>R f x) has_vector_derivative (c *\<^sub>R f')) net" | |
| 44140 
2c10c35dd4be
remove several redundant and unused theorems about derivatives
 huffman parents: 
44137diff
changeset | 2191 | unfolding has_vector_derivative_def | 
| 44282 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44140diff
changeset | 2192 | apply (drule scaleR_right_has_derivative) | 
| 53781 | 2193 | apply (auto simp add: algebra_simps) | 
| 2194 | done | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2195 | |
| 44123 | 2196 | lemma has_vector_derivative_cmul_eq: | 
| 2197 | assumes "c \<noteq> 0" | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2198 | shows "(((\<lambda>x. c *\<^sub>R f x) has_vector_derivative (c *\<^sub>R f')) net \<longleftrightarrow> (f has_vector_derivative f') net)" | 
| 53781 | 2199 | apply rule | 
| 2200 | apply (drule has_vector_derivative_cmul[where c="1/c"]) | |
| 2201 | defer | |
| 2202 | apply (rule has_vector_derivative_cmul) | |
| 2203 | using assms | |
| 2204 | apply auto | |
| 2205 | done | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2206 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2207 | lemma has_vector_derivative_neg: | 
| 53781 | 2208 | "(f has_vector_derivative f') net \<Longrightarrow> ((\<lambda>x. - f x) has_vector_derivative (- f')) net" | 
| 2209 | unfolding has_vector_derivative_def | |
| 2210 | apply (drule has_derivative_neg) | |
| 2211 | apply auto | |
| 2212 | done | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2213 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2214 | lemma has_vector_derivative_add: | 
| 44123 | 2215 | assumes "(f has_vector_derivative f') net" | 
| 53781 | 2216 | and "(g has_vector_derivative g') net" | 
| 2217 | shows "((\<lambda>x. f x + g x) has_vector_derivative (f' + g')) net" | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2218 | using has_derivative_add[OF assms[unfolded has_vector_derivative_def]] | 
| 53781 | 2219 | unfolding has_vector_derivative_def | 
| 2220 | unfolding scaleR_right_distrib | |
| 2221 | 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 | 2222 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2223 | lemma has_vector_derivative_sub: | 
| 44123 | 2224 | assumes "(f has_vector_derivative f') net" | 
| 53781 | 2225 | and "(g has_vector_derivative g') net" | 
| 2226 | shows "((\<lambda>x. f x - g x) has_vector_derivative (f' - g')) net" | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2227 | using has_derivative_sub[OF assms[unfolded has_vector_derivative_def]] | 
| 53781 | 2228 | unfolding has_vector_derivative_def scaleR_right_diff_distrib | 
| 2229 | 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 | 2230 | |
| 37650 | 2231 | lemma has_vector_derivative_bilinear_within: | 
| 44123 | 2232 | assumes "(f has_vector_derivative f') (at x within s)" | 
| 53781 | 2233 | and "(g has_vector_derivative g') (at x within s)" | 
| 44123 | 2234 | assumes "bounded_bilinear h" | 
| 2235 | shows "((\<lambda>x. h (f x) (g x)) has_vector_derivative (h (f x) g' + h f' (g x))) (at x within s)" | |
| 53781 | 2236 | proof - | 
| 2237 | interpret bounded_bilinear h | |
| 2238 | using assms by auto | |
| 2239 | show ?thesis | |
| 2240 | using has_derivative_bilinear_within[OF assms(1-2)[unfolded has_vector_derivative_def], of h] | |
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36844diff
changeset | 2241 | unfolding o_def has_vector_derivative_def | 
| 53781 | 2242 | using assms(3) | 
| 2243 | unfolding scaleR_right scaleR_left scaleR_right_distrib | |
| 44123 | 2244 | by auto | 
| 2245 | qed | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2246 | |
| 37650 | 2247 | lemma has_vector_derivative_bilinear_at: | 
| 44123 | 2248 | assumes "(f has_vector_derivative f') (at x)" | 
| 53781 | 2249 | and "(g has_vector_derivative g') (at x)" | 
| 44123 | 2250 | assumes "bounded_bilinear h" | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2251 | shows "((\<lambda>x. h (f x) (g x)) has_vector_derivative (h (f x) g' + h f' (g x))) (at x)" | 
| 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 | 2252 | using has_vector_derivative_bilinear_within[OF assms] . | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2253 | |
| 44123 | 2254 | lemma has_vector_derivative_at_within: | 
| 2255 | "(f has_vector_derivative f') (at x) \<Longrightarrow> (f has_vector_derivative f') (at x within s)" | |
| 2256 | unfolding has_vector_derivative_def | |
| 45031 | 2257 | 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 | 2258 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2259 | lemma has_vector_derivative_transform_within: | 
| 53781 | 2260 | assumes "0 < d" | 
| 2261 | and "x \<in> s" | |
| 2262 | and "\<forall>x'\<in>s. dist x' x < d \<longrightarrow> f x' = g x'" | |
| 44123 | 2263 | assumes "(f has_vector_derivative f') (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 | 2264 | shows "(g has_vector_derivative f') (at x within s)" | 
| 53781 | 2265 | using assms | 
| 2266 | unfolding has_vector_derivative_def | |
| 44123 | 2267 | 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 | 2268 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2269 | lemma has_vector_derivative_transform_at: | 
| 53781 | 2270 | assumes "0 < d" | 
| 2271 | and "\<forall>x'. dist x' x < d \<longrightarrow> f x' = g x'" | |
| 2272 | and "(f has_vector_derivative f') (at x)" | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2273 | shows "(g has_vector_derivative f') (at x)" | 
| 53781 | 2274 | using assms | 
| 2275 | unfolding has_vector_derivative_def | |
| 44123 | 2276 | by (rule has_derivative_transform_at) | 
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2277 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2278 | lemma has_vector_derivative_transform_within_open: | 
| 53781 | 2279 | assumes "open s" | 
| 2280 | and "x \<in> s" | |
| 2281 | and "\<forall>y\<in>s. f y = g y" | |
| 2282 | and "(f has_vector_derivative f') (at x)" | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2283 | shows "(g has_vector_derivative f') (at x)" | 
| 53781 | 2284 | using assms | 
| 2285 | unfolding has_vector_derivative_def | |
| 44123 | 2286 | 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 | 2287 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2288 | lemma vector_diff_chain_at: | 
| 44123 | 2289 | assumes "(f has_vector_derivative f') (at x)" | 
| 53781 | 2290 | 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 | 2291 | shows "((g \<circ> f) has_vector_derivative (f' *\<^sub>R g')) (at x)" | 
| 53781 | 2292 | using assms(2) | 
| 2293 | unfolding has_vector_derivative_def | |
| 2294 | apply - | |
| 2295 | apply (drule diff_chain_at[OF assms(1)[unfolded has_vector_derivative_def]]) | |
| 2296 | apply (simp only: o_def real_scaleR_def scaleR_scaleR) | |
| 2297 | done | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2298 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2299 | lemma vector_diff_chain_within: | 
| 44123 | 2300 | assumes "(f has_vector_derivative f') (at x within s)" | 
| 53781 | 2301 | and "(g has_vector_derivative g') (at (f x) within f ` s)" | 
| 2302 | shows "((g \<circ> f) has_vector_derivative (f' *\<^sub>R g')) (at x within s)" | |
| 2303 | using assms(2) | |
| 2304 | unfolding has_vector_derivative_def | |
| 2305 | apply - | |
| 2306 | apply (drule diff_chain_within[OF assms(1)[unfolded has_vector_derivative_def]]) | |
| 2307 | apply (simp only: o_def real_scaleR_def scaleR_scaleR) | |
| 2308 | done | |
| 33741 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2309 | |
| 
4c414d0835ab
Added derivation and Brouwer's fixpoint theorem in Multivariate Analysis (translated by Robert Himmelmann from HOL-light)
 hoelzl parents: diff
changeset | 2310 | end |