diff --git a/thys/Fourier/Fourier.thy b/thys/Fourier/Fourier.thy --- a/thys/Fourier/Fourier.thy +++ b/thys/Fourier/Fourier.thy @@ -1,2742 +1,2740 @@ section\The basics of Fourier series\ text\Ported from HOL Light; thanks to Manuel Eberl for help with the real asymp proof methods\ theory "Fourier" imports Periodic Square_Integrable "HOL-Real_Asymp.Real_Asymp" Confine Fourier_Aux2 begin subsection\Orthonormal system of L2 functions and their Fourier coefficients\ definition orthonormal_system :: "'a::euclidean_space set \ ('b \ 'a \ real) \ bool" where "orthonormal_system S w \ \m n. l2product S (w m) (w n) = (if m = n then 1 else 0)" definition orthonormal_coeff :: "'a::euclidean_space set \ (nat \ 'a \ real) \ ('a \ real) \ nat \ real" where "orthonormal_coeff S w f n = l2product S (w n) f" lemma orthonormal_system_eq: "orthonormal_system S w \ l2product S (w m) (w n) = (if m = n then 1 else 0)" by (simp add: orthonormal_system_def) lemma orthonormal_system_l2norm: "orthonormal_system S w \ l2norm S (w i) = 1" by (simp add: l2norm_def orthonormal_system_def) lemma orthonormal_partial_sum_diff: assumes os: "orthonormal_system S w" and w: "\i. (w i) square_integrable S" and f: "f square_integrable S" and "finite I" shows "(l2norm S (\x. f x - (\i\I. a i * w i x)))\<^sup>2 = (l2norm S f)\<^sup>2 + (\i\I. (a i)\<^sup>2) - 2 * (\i\I. a i * orthonormal_coeff S w f i)" proof - have "S \ sets lebesgue" using f square_integrable_def by blast then have "(\x. \i\I. a i * w i x) square_integrable S" by (intro square_integrable_sum square_integrable_lmult w \finite I\) with assms show ?thesis apply (simp add: l2norm_pow_2 orthonormal_coeff_def orthonormal_system_def) apply (simp add: l2product_rdiff l2product_sym l2product_rsum l2product_rmult algebra_simps power2_eq_square if_distrib [of "\x. _ * x"]) done qed lemma orthonormal_optimal_partial_sum: assumes "orthonormal_system S w" "\i. (w i) square_integrable S" "f square_integrable S" "finite I" shows "l2norm S (\x. f x - (\i\I. orthonormal_coeff S w f i * w i x)) \ l2norm S (\x. f x - (\i\I. a i * w i x))" proof - have "2 * (a i * l2product S f(w i)) \ (a i)\<^sup>2 + (l2product S f(w i))\<^sup>2" for i using sum_squares_bound [of "a i" "l2product S f(w i)"] by (force simp: algebra_simps) then have *: "2 * (\i\I. a i * l2product S f(w i)) \ (\i\I. (a i)\<^sup>2 + (l2product S f(w i))\<^sup>2)" by (simp add: sum_distrib_left sum_mono) have S: "S \ sets lebesgue" using assms square_integrable_def by blast with assms * show ?thesis apply (simp add: l2norm_le square_integrable_sum square_integrable_diff square_integrable_lmult flip: l2norm_pow_2) apply (simp add: orthonormal_coeff_def orthonormal_partial_sum_diff) apply (simp add: l2product_sym power2_eq_square sum.distrib) done qed lemma Bessel_inequality: assumes "orthonormal_system S w" "\i. (w i) square_integrable S" "f square_integrable S" "finite I" shows "(\i\I. (orthonormal_coeff S w f i)\<^sup>2) \ (l2norm S f)\<^sup>2" using orthonormal_partial_sum_diff [OF assms, of "orthonormal_coeff S w f"] apply (simp add: algebra_simps flip: power2_eq_square) by (metis (lifting) add_diff_cancel_left' diff_ge_0_iff_ge zero_le_power2) lemma Fourier_series_square_summable: assumes os: "orthonormal_system S w" and w: "\i. (w i) square_integrable S" and f: "f square_integrable S" shows "summable (confine (\i. (orthonormal_coeff S w f i) ^ 2) I)" proof - have "incseq (\n. \i\I \ {..n}. (orthonormal_coeff S w f i)\<^sup>2)" by (auto simp: incseq_def intro: sum_mono2) moreover have "\i. (\i\I \ {..i}. (orthonormal_coeff S w f i)\<^sup>2) \ (l2norm S f)\<^sup>2" by (simp add: Bessel_inequality assms) ultimately obtain L where "(\n. \i\I \ {..n}. (orthonormal_coeff S w f i)\<^sup>2) \ L" using incseq_convergent by blast then have "\r>0. \no. \n\no. \(\i\{..n} \ I. (orthonormal_coeff S w f i)\<^sup>2) - L\ < r" by (simp add: LIMSEQ_iff Int_commute) then show ?thesis by (auto simp: summable_def sums_confine_le LIMSEQ_iff) qed lemma orthonormal_Fourier_partial_sum_diff_squared: assumes os: "orthonormal_system S w" and w: "\i. (w i) square_integrable S" and f: "f square_integrable S" and "finite I" shows "(l2norm S (\x. f x -(\i\I. orthonormal_coeff S w f i * w i x)))\<^sup>2 = (l2norm S f)\<^sup>2 - (\i\I. (orthonormal_coeff S w f i)\<^sup>2)" using orthonormal_partial_sum_diff [OF assms, where a = "orthonormal_coeff S w f"] by (simp add: power2_eq_square) lemma Fourier_series_l2_summable: assumes os: "orthonormal_system S w" and w: "\i. (w i) square_integrable S" and f: "f square_integrable S" obtains g where "g square_integrable S" "(\n. l2norm S (\x. (\i\I \ {..n}. orthonormal_coeff S w f i * w i x) - g x)) \ 0" proof - have S: "S \ sets lebesgue" using f square_integrable_def by blast let ?\ = "\A x. (\i\I \ A. orthonormal_coeff S w f i * w i x)" let ?\ = "confine (\i. (orthonormal_coeff S w f i)\<^sup>2) I" have si: "?\ A square_integrable S" if "finite A" for A by (force intro: square_integrable_lmult w square_integrable_sum S that) have "\N. \m\N. \n\N. l2norm S (\x. ?\ {..m} x - ?\ {..n} x) < e" if "e > 0" for e proof - have "summable ?\" by (rule Fourier_series_square_summable [OF os w f]) then have "Cauchy (\n. sum ?\ {..n})" by (simp add: summable_def sums_def_le convergent_eq_Cauchy) then obtain M where M: "\m n. \m\M; n\M\ \ dist (sum ?\ {..m}) (sum ?\ {..n}) < e\<^sup>2" using \e > 0\ unfolding Cauchy_def by (drule_tac x="e^2" in spec) auto have "\m\M; n\M\ \ l2norm S (\x. ?\ {..m} x - ?\ {..n} x) < e" for m n proof (induct m n rule: linorder_class.linorder_wlog) case (le m n) have sum_diff: "sum f(I \ {..n}) - sum f(I \ {..m}) = sum f(I \ {Suc m..n})" for f :: "nat \ real" proof - have "(I \ {..n}) = (I \ {..m} \ I \ {Suc m..n})" "(I \ {..m}) \ (I \ {Suc m..n}) = {}" using le by auto then show ?thesis by (simp add: algebra_simps flip: sum.union_disjoint) qed have "(l2norm S (\x. ?\ {..n} x - ?\ {..m} x))^2 \ \(\i\I \ {..n}. (orthonormal_coeff S w f i)\<^sup>2) - (\i\I \ {..m}. (orthonormal_coeff S w f i)\<^sup>2)\" proof (simp add: sum_diff) have "(l2norm S (?\ {Suc m..n}))\<^sup>2 = (\i\I \ {Suc m..n}. l2product S (\x. \j\I \ {Suc m..n}. orthonormal_coeff S w f j * w j x) (\x. orthonormal_coeff S w f i * w i x))" by (simp add: l2norm_pow_2 l2product_rsum si w) also have "\ = (\i\I \ {Suc m..n}. \j\I \ {Suc m..n}. orthonormal_coeff S w f j * orthonormal_coeff S w f i * l2product S (w j) (w i))" by (simp add: l2product_lsum l2product_lmult l2product_rmult si w flip: mult.assoc) also have "\ \ \\i\I \ {Suc m..n}. (orthonormal_coeff S w f i)\<^sup>2\" by (simp add: orthonormal_system_eq [OF os] power2_eq_square if_distrib [of "(*)_"] cong: if_cong) finally show "(l2norm S (?\ {Suc m..n}))\<^sup>2 \ \\i\I \ {Suc m..n}. (orthonormal_coeff S w f i)\<^sup>2\" . qed also have "\ < e\<^sup>2" using M [OF \n\M\ \m\M\] by (simp add: dist_real_def sum_confine_eq_Int Int_commute) finally have "(l2norm S (\x. ?\ {..m} x - ?\ {..n} x))^2 < e^2" using l2norm_diff [OF si si] by simp with \e > 0\ show ?case by (simp add: power2_less_imp_less) next case (sym a b) then show ?case by (subst l2norm_diff) (auto simp: si) qed then show ?thesis by blast qed with that show thesis by (blast intro: si [of "{.._}"] l2_complete [of "\n. ?\ {..n}"]) qed lemma Fourier_series_l2_summable_strong: assumes os: "orthonormal_system S w" and w: "\i. (w i) square_integrable S" and f: "f square_integrable S" obtains g where "g square_integrable S" "\i. i \ I \ orthonormal_coeff S w (\x. f x - g x) i = 0" "(\n. l2norm S (\x. (\i\I \ {..n}. orthonormal_coeff S w f i * w i x) - g x)) \ 0" proof - have S: "S \ sets lebesgue" using f square_integrable_def by blast obtain g where g: "g square_integrable S" and teng: "(\n. l2norm S (\x. (\i\I \ {..n}. orthonormal_coeff S w f i * w i x) - g x)) \ 0" using Fourier_series_l2_summable [OF assms] . show thesis proof show "orthonormal_coeff S w (\x. f x - g x) i = 0" if "i \ I" for i proof (rule tendsto_unique) let ?\ = "\A x. (\i\I \ A. orthonormal_coeff S w f i * w i x)" let ?f = "\n. l2product S (w i) (\x. (f x - ?\ {..n} x) + (?\ {..n} x - g x))" show "?f \ orthonormal_coeff S w (\x. f x - g x) i" by (simp add: orthonormal_coeff_def) have 1: "(\n. l2product S (w i) (\x. f x - ?\ {..n} x)) \ 0" proof (rule tendsto_eventually) have "l2product S (w i) (\x. f x - ?\ {..j} x) = 0" if "j \ i" for j using that \i \ I\ apply (simp add: l2product_rdiff l2product_rsum l2product_rmult orthonormal_coeff_def w f S) apply (simp add: orthonormal_system_eq [OF os] if_distrib [of "(*)_"] cong: if_cong) done then show "\\<^sub>F n in sequentially. l2product S (w i) (\x. f x - ?\ {..n} x) = 0" using eventually_at_top_linorder by blast qed have 2: "(\n. l2product S (w i) (\x. ?\ {..n} x - g x)) \ 0" proof (intro Lim_null_comparison [OF _ teng] eventuallyI) show "norm (l2product S (w i) (\x. ?\ {..n} x - g x)) \ l2norm S (\x. ?\ {..n} x - g x)" for n using Schwartz_inequality_abs [of "w i" S "(\x. ?\ {..n} x - g x)"] by (simp add: S g f w orthonormal_system_l2norm [OF os]) qed show "?f \ 0" using that tendsto_add [OF 1 2] by (subst l2product_radd) (simp_all add: l2product_radd w f g S) qed auto qed (auto simp: g teng) qed subsection\Actual trigonometric orthogonality relations\ lemma integrable_sin_cx: "integrable (lebesgue_on {-pi..pi}) (\x. sin(x * c))" by (intro continuous_imp_integrable_real continuous_intros) lemma integrable_cos_cx: "integrable (lebesgue_on {-pi..pi}) (\x. cos(x * c))" by (intro continuous_imp_integrable_real continuous_intros) lemma integral_cos_Z' [simp]: assumes "n \ \" shows "integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. cos(n * x)) = (if n = 0 then 2 * pi else 0)" by (metis assms integral_cos_Z mult_commute_abs) lemma integral_sin_and_cos: fixes m n::int shows "integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. cos(m * x) * cos(n * x)) = (if \m\ = \n\ then if n = 0 then 2 * pi else pi else 0)" "integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. cos(m * x) * sin(n * x)) = 0" "integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. sin(m * x) * cos(n * x)) = 0" "\m \ 0; n \ 0\ \ integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. sin (m * x) * sin (n * x)) = (if m = n \ n \ 0 then pi else 0)" "\integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. sin (m * x) * sin (n * x))\ = (if \m\ = \n\ \ n \ 0 then pi else 0)" by (simp_all add: abs_if sin_times_sin cos_times_sin sin_times_cos cos_times_cos integrable_sin_cx integrable_cos_cx mult_ac flip: distrib_left distrib_right left_diff_distrib right_diff_distrib) lemma integral_sin_and_cos_Z [simp]: fixes m n::real assumes "m \ \" "n \ \" shows "integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. cos(m * x) * cos(n * x)) = (if \m\ = \n\ then if n = 0 then 2 * pi else pi else 0)" "integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. cos(m * x) * sin(n * x)) = 0" "integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. sin(m * x) * cos(n * x)) = 0" "\integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. sin (m * x) * sin (n * x))\ = (if \m\ = \n\ \ n \ 0 then pi else 0)" using assms unfolding Ints_def apply safe unfolding integral_sin_and_cos apply auto done lemma integral_sin_and_cos_N [simp]: fixes m n::real assumes "m \ \" "n \ \" shows "integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. sin (m * x) * sin (n * x)) = (if m = n \ n \ 0 then pi else 0)" using assms unfolding Nats_altdef1 by (auto simp: integral_sin_and_cos) lemma integrable_sin_and_cos: fixes m n::int shows "integrable (lebesgue_on {a..b}) (\x. cos(x * m) * cos(x * n))" "integrable (lebesgue_on {a..b}) (\x. cos(x * m) * sin(x * n))" "integrable (lebesgue_on {a..b}) (\x. sin(x * m) * cos(x * n))" "integrable (lebesgue_on {a..b}) (\x. sin(x * m) * sin(x * n))" by (intro continuous_imp_integrable_real continuous_intros)+ lemma sqrt_pi_ge1: "sqrt pi \ 1" using pi_gt3 by auto definition trigonometric_set :: "nat \ real \ real" where "trigonometric_set n \ if n = 0 then \x. 1 / sqrt(2 * pi) else if odd n then \x. sin(real(Suc (n div 2)) * x) / sqrt(pi) else (\x. cos((n div 2) * x) / sqrt pi)" lemma trigonometric_set: "trigonometric_set 0 x = 1 / sqrt(2 * pi)" "trigonometric_set (Suc (2 * n)) x = sin(real(Suc n) * x) / sqrt(pi)" "trigonometric_set (2 * n + 2) x = cos(real(Suc n) * x) / sqrt(pi)" "trigonometric_set (Suc (Suc (2 * n))) x = cos(real(Suc n) * x) / sqrt(pi)" by (simp_all add: trigonometric_set_def algebra_simps add_divide_distrib) lemma trigonometric_set_even: "trigonometric_set(2*k) = (if k = 0 then (\x. 1 / sqrt(2 * pi)) else (\x. cos(k * x) / sqrt pi))" by (induction k) (auto simp: trigonometric_set_def add_divide_distrib split: if_split_asm) lemma orthonormal_system_trigonometric_set: "orthonormal_system {-pi..pi} trigonometric_set" proof - have "l2product {-pi..pi} (trigonometric_set m) (trigonometric_set n) = (if m = n then 1 else 0)" for m n proof (induction m rule: odd_even_cases) case 0 show ?case by (induction n rule: odd_even_cases) (auto simp: trigonometric_set l2product_def measure_restrict_space) next case (odd m) show ?case by (induction n rule: odd_even_cases) (auto simp: trigonometric_set l2product_def double_not_eq_Suc_double) next case (even m) show ?case by (induction n rule: odd_even_cases) (auto simp: trigonometric_set l2product_def Suc_double_not_eq_double) qed then show ?thesis unfolding orthonormal_system_def by auto qed lemma square_integrable_trigonometric_set: "(trigonometric_set i) square_integrable {-pi..pi}" proof - have "continuous_on {-pi..pi} (\x. sin ((1 + real n) * x) / sqrt pi)" for n by (intro continuous_intros) auto moreover have "continuous_on {-pi..pi} (\x. cos (real i * x / 2) / sqrt pi) " by (intro continuous_intros) auto ultimately show ?thesis by (simp add: trigonometric_set_def) qed subsection\Weierstrass for trigonometric polynomials\ lemma Weierstrass_trig_1: fixes g :: "real \ real" assumes contf: "continuous_on UNIV g" and periodic: "\x. g(x + 2 * pi) = g x" and 1: "norm z = 1" shows "continuous (at z within (sphere 0 1)) (g \ Im \ Ln)" proof (cases "Re z < 0") let ?f = "complex_of_real \ g \ (\x. x + pi) \ Im \ Ln \ uminus" case True have "continuous (at z within (sphere 0 1)) (complex_of_real \ g \ Im \ Ln)" proof (rule continuous_transform_within) have [simp]: "z \ \\<^sub>\\<^sub>0" using True complex_nonneg_Reals_iff by auto show "continuous (at z within (sphere 0 1)) ?f" by (intro continuous_within_Ln continuous_intros continuous_on_imp_continuous_within [OF contf]) auto show "?f x' = (complex_of_real \ g \ Im \ Ln) x'" if "x' \ (sphere 0 1)" and "dist x' z < 1" for x' proof - have "x' \ 0" using that by auto with that show ?thesis by (auto simp: Ln_minus add.commute periodic) qed qed (use 1 in auto) then have "continuous (at z within (sphere 0 1)) (Re \ complex_of_real \ g \ Im \ Ln)" unfolding o_def by (rule continuous_Re) then show ?thesis by (simp add: o_def) next case False let ?f = "complex_of_real \ g \ Im \ Ln \ uminus" have "z \ 0" using 1 by auto with False have "z \ \\<^sub>\\<^sub>0" by (auto simp: not_less nonpos_Reals_def) then have "continuous (at z within (sphere 0 1)) (complex_of_real \ g \ Im \ Ln)" by (intro continuous_within_Ln continuous_intros continuous_on_imp_continuous_within [OF contf]) auto then have "continuous (at z within (sphere 0 1)) (Re \ complex_of_real \ g \ Im \ Ln)" unfolding o_def by (rule continuous_Re) then show ?thesis by (simp add: o_def) qed inductive_set cx_poly :: "(complex \ real) set" where Re: "Re \ cx_poly" | Im: "Im \ cx_poly" | const: "(\x. c) \ cx_poly" | add: "\f \ cx_poly; g \ cx_poly\ \ (\x. f x + g x) \ cx_poly" | mult: "\f \ cx_poly; g \ cx_poly\ \ (\x. f x * g x) \ cx_poly" declare cx_poly.intros [intro] lemma Weierstrass_trig_polynomial: assumes contf: "continuous_on {-pi..pi} f" and fpi: "f(-pi) = f pi" and "0 < e" obtains n::nat and a b where "\x::real. x \ {-pi..pi} \ \f x - (\k\n. a k * sin (k * x) + b k * cos (k * x))\ < e" proof - interpret CR: function_ring_on "cx_poly" "sphere 0 1" proof show "continuous_on (sphere 0 1) f" if "f \ cx_poly" for f using that by induction (assumption | intro continuous_intros)+ fix x y::complex assume "x \ sphere 0 1" and "y \ sphere 0 1" and "x \ y" then consider "Re x \ Re y" | "Im x \ Im y" using complex_eqI by blast then show "\f\cx_poly. f x \ f y" by (metis cx_poly.Im cx_poly.Re) qed (auto simp: cx_poly.intros) have "continuous (at z within (sphere 0 1)) (f \ Im \ Ln)" if "norm z = 1" for z proof - obtain g where contg: "continuous_on UNIV g" and gf: "\x. x \ {-pi..pi} \ g x = f x" and periodic: "\x. g(x + 2*pi) = g x" using Tietze_periodic_interval [OF contf fpi] by auto let ?f = "(g \ Im \ Ln)" show ?thesis proof (rule continuous_transform_within) show "continuous (at z within (sphere 0 1)) ?f" using Weierstrass_trig_1 [OF contg periodic that] by (simp add: sphere_def) show "?f x' = (f \ Im \ Ln) x'" if "x' \ sphere 0 1" "dist x' z < 1" for x' proof - have "x' \ 0" using that by auto then have "Im (Ln x') \ {-pi..pi}" using Im_Ln_le_pi [of x'] mpi_less_Im_Ln [of x'] by simp then show ?thesis using gf by simp qed qed (use that in auto) qed then have "continuous_on (sphere 0 1) (f \ Im \ Ln)" using continuous_on_eq_continuous_within mem_sphere_0 by blast then obtain g where "g \ cx_poly" and g: "\x. x \ sphere 0 1 \ \(f \ Im \ Ln) x - g x\ < e" using CR.Stone_Weierstrass_basic [of "f \ Im \ Ln"] \e > 0\ by meson define trigpoly where "trigpoly \ \f. \n a b. f = (\x. (\k\n. a k * sin(real k * x) + b k * cos(real k * x)))" have tp_const: "trigpoly(\x. c)" for c unfolding trigpoly_def by (rule_tac x=0 in exI) auto have tp_add: "trigpoly(\x. f x + g x)" if "trigpoly f" "trigpoly g" for f g proof - obtain n1 a1 b1 where feq: "f = (\x. \k\n1. a1 k * sin (real k * x) + b1 k * cos (real k * x))" using \trigpoly f\ trigpoly_def by blast obtain n2 a2 b2 where geq: "g = (\x. \k\n2. a2 k * sin (real k * x) + b2 k * cos (real k * x))" using \trigpoly g\ trigpoly_def by blast let ?a = "\n. (if n \ n1 then a1 n else 0) + (if n \ n2 then a2 n else 0)" let ?b = "\n. (if n \ n1 then b1 n else 0) + (if n \ n2 then b2 n else 0)" have eq: "{k. k \ max a b \ k \ a} = {..a}" "{k. k \ max a b \ k \ b} = {..b}" for a b::nat by auto have "(\x. f x + g x) = (\x. \k \ max n1 n2. ?a k * sin (real k * x) + ?b k * cos (real k * x))" by (simp add: feq geq algebra_simps eq sum.distrib if_distrib [of "\u. _ * u"] cong: if_cong flip: sum.inter_filter) then show ?thesis unfolding trigpoly_def by meson qed have tp_sum: "trigpoly(\x. \i\S. f i x)" if "finite S" "\i. i \ S \ trigpoly(f i)" for f and S :: "nat set" using that by induction (auto simp: tp_const tp_add) have tp_cmul: "trigpoly(\x. c * f x)" if "trigpoly f" for f c proof - obtain n a b where feq: "f = (\x. \k\n. a k * sin (real k * x) + b k * cos (real k * x))" using \trigpoly f\ trigpoly_def by blast have "(\x. c * f x) = (\x. \k \ n. (c * a k) * sin (real k * x) + (c * b k) * cos (real k * x))" by (simp add: feq algebra_simps sum_distrib_left) then show ?thesis unfolding trigpoly_def by meson qed have tp_cdiv: "trigpoly(\x. f x / c)" if "trigpoly f" for f c using tp_cmul [OF \trigpoly f\, of "inverse c"] by (simp add: divide_inverse_commute) have tp_diff: "trigpoly(\x. f x - g x)" if "trigpoly f" "trigpoly g" for f g using tp_add [OF \trigpoly f\ tp_cmul [OF \trigpoly g\, of "-1"]] by auto have tp_sin: "trigpoly(\x. sin (n * x))" if "n \ \" for n unfolding trigpoly_def proof obtain k where "n = real k" using Nats_cases \n \ \\ by blast then show "\a b. (\x. sin (n * x)) = (\x. \i \ nat\n\. a i * sin (real i * x) + b i * cos (real i * x))" apply (rule_tac x="\i. if i = k then 1 else 0" in exI) apply (rule_tac x="\i. 0" in exI) apply (simp add: if_distrib [of "\u. u * _"] cong: if_cong) done qed have tp_cos: "trigpoly(\x. cos (n * x))" if "n \ \" for n unfolding trigpoly_def proof obtain k where "n = real k" using Nats_cases \n \ \\ by blast then show "\a b. (\x. cos (n * x)) = (\x. \i \ nat\n\. a i * sin (real i * x) + b i * cos (real i * x))" apply (rule_tac x="\i. 0" in exI) apply (rule_tac x="\i. if i = k then 1 else 0" in exI) apply (simp add: if_distrib [of "\u. u * _"] cong: if_cong) done qed have tp_sincos: "trigpoly(\x. sin(i * x) * sin(j * x)) \ trigpoly(\x. sin(i * x) * cos(j * x)) \ trigpoly(\x. cos(i * x) * sin(j * x)) \ trigpoly(\x. cos(i * x) * cos(j * x))" (is "?P i j") for i j::nat proof (rule linorder_wlog [of _ j i]) show "?P j i" if "i \ j" for j i using that by (simp add: sin_times_sin cos_times_cos sin_times_cos cos_times_sin diff_divide_distrib tp_add tp_diff tp_cdiv tp_cos tp_sin flip: left_diff_distrib distrib_right) qed (simp add: mult_ac) have tp_mult: "trigpoly(\x. f x * g x)" if "trigpoly f" "trigpoly g" for f g proof - obtain n1 a1 b1 where feq: "f = (\x. \k\n1. a1 k * sin (real k * x) + b1 k * cos (real k * x))" using \trigpoly f\ trigpoly_def by blast obtain n2 a2 b2 where geq: "g = (\x. \k\n2. a2 k * sin (real k * x) + b2 k * cos (real k * x))" using \trigpoly g\ trigpoly_def by blast then show ?thesis unfolding feq geq by (simp add: algebra_simps sum_product tp_sum tp_add tp_cmul tp_sincos del: mult.commute) qed have *: "trigpoly(\x. f(exp(\ * of_real x)))" if "f \ cx_poly" for f using that proof induction case Re then show ?case using tp_cos [of 1] by (auto simp: Re_exp) next case Im then show ?case using tp_sin [of 1] by (auto simp: Im_exp) qed (auto simp: tp_const tp_add tp_mult) obtain n a b where eq: "(g (iexp x)) = (\k\n. a k * sin (real k * x) + b k * cos (real k * x))" for x using * [OF \g \ cx_poly\] trigpoly_def by meson show thesis proof show "\f \ - (\k\n. a k * sin (real k * \) + b k * cos (real k * \))\ < e" if "\ \ {-pi..pi}" for \ proof - have "f \ - g (iexp \) = (f \ Im \ Ln) (iexp \) - g (iexp \)" proof (cases "\ = -pi") case True then show ?thesis by (simp add: exp_minus fpi) next case False then show ?thesis using that by auto qed then show ?thesis using g [of "exp(\ * of_real \)"] by (simp flip: eq) qed qed qed subsection\A bit of extra hacking round so that the ends of a function are OK\ lemma integral_tweak_ends: fixes a b :: real assumes "a < b" "e > 0" obtains f where "continuous_on {a..b} f" "f a = d" "f b = 0" "l2norm {a..b} f < e" proof - have cont: "continuous_on {a..b} (\x. if x \ a+1 / real(Suc n) then ((Suc n) * d) * ((a+1 / (Suc n)) - x) else 0)" for n proof (cases "a+1/(Suc n) \ b") case True have *: "1 / (1 + real n) > 0" by auto have abeq: "{a..b} = {a..a+1/(Suc n)} \ {a+1/(Suc n)..b}" using \a < b\ True apply auto using "*" by linarith show ?thesis unfolding abeq proof (rule continuous_on_cases) show "continuous_on {a..a+1 / real (Suc n)} (\x. real (Suc n) * d * (a+1 / real (Suc n) - x))" by (intro continuous_intros) qed auto next case False show ?thesis proof (rule continuous_on_eq [where f = "\x. ((Suc n) * d) * ((a+1/(Suc n)) - x)"]) show " continuous_on {a..b} (\x. (Suc n) * d * (a+1 / real (Suc n) - x))" by (intro continuous_intros) qed (use False in auto) qed let ?f = "\k x. (if x \ a+1 / (Suc k) then (Suc k) * d * (a+1 / (Suc k) - x) else 0)\<^sup>2" let ?g = "\x. if x = a then d\<^sup>2 else 0" have bmg: "?g \ borel_measurable (lebesgue_on {a..b})" apply (rule measurable_restrict_space1) using borel_measurable_if_I [of _ "{a}", OF borel_measurable_const] by auto have bmf: "?f k \ borel_measurable (lebesgue_on {a..b})" for k proof - have bm: "(\x. (Suc k) * d * (a+1 / real (Suc k) - x)) \ borel_measurable (lebesgue_on {..a+1 / (Suc k)})" by (intro measurable) (auto simp: measurable_completion measurable_restrict_space1) show ?thesis apply (intro borel_measurable_power measurable_restrict_space1) using borel_measurable_if_I [of _ "{.. a+1 / (Suc k)}", OF bm] apply auto done qed have int_d2: "integrable (lebesgue_on {a..b}) (\x. d\<^sup>2)" by (intro continuous_imp_integrable_real continuous_intros) have "(\k. ?f k x) \ ?g x" if x: "x \ {a..b}" for x proof (cases "x = a") case False then have "x > a" using x by auto with x obtain N where "N > 0" and N: "1 / real N < x-a" using real_arch_invD [of "x-a"] by (force simp: inverse_eq_divide) then have "x > a+1 / (1 + real k)" if "k \ N" for k proof - have "a+1 / (1 + real k) < a+1 / real N" using that \0 < N\ by (simp add: field_simps) also have "\ < x" using N by linarith finally show ?thesis . qed then show ?thesis apply (intro tendsto_eventually eventually_sequentiallyI [where c=N]) by (fastforce simp: False) qed auto then have tends: "AE x in (lebesgue_on {a..b}). (\k. ?f k x) \ ?g x" by force have le_d2: "\k. AE x in (lebesgue_on {a..b}). norm (?f k x) \ d\<^sup>2" proof show "norm ((if x \ a+1 / real (Suc k) then real (Suc k) * d * (a+1 / real (Suc k) - x) else 0)\<^sup>2) \ d\<^sup>2" if "x \ space (lebesgue_on {a..b})" for k x using that apply (simp add: abs_mult divide_simps flip: abs_le_square_iff) apply (auto simp: abs_if zero_less_mult_iff mult_left_le) done qed have integ: "integrable (lebesgue_on {a..b}) ?g" using integrable_dominated_convergence [OF bmg bmf int_d2 tends le_d2] by auto have int: "(\k. integral\<^sup>L (lebesgue_on {a..b}) (?f k)) \ integral\<^sup>L (lebesgue_on {a..b}) ?g" using integral_dominated_convergence [OF bmg bmf int_d2 tends le_d2] by auto then obtain N where N: "\k. k \ N \ \integral\<^sup>L (lebesgue_on {a..b}) (?f k) - integral\<^sup>L (lebesgue_on {a..b}) ?g\ < e\<^sup>2" apply (simp add: lim_sequentially dist_real_def) apply (drule_tac x="e^2" in spec) using \e > 0\ by auto obtain M where "M > 0" and M: "1 / real M < b-a" using real_arch_invD [of "b-a"] by (metis \a < b\ diff_gt_0_iff_gt inverse_eq_divide neq0_conv) have *: "\integral\<^sup>L (lebesgue_on {a..b}) (?f (max M N)) - integral\<^sup>L (lebesgue_on {a..b}) ?g\ < e\<^sup>2" using N by force let ?\ = "\x. if x \ a+1/(Suc (max M N)) then ((Suc (max M N)) * d) * ((a+1/(Suc (max M N))) - x) else 0" show ?thesis proof show "continuous_on {a..b} ?\" by (rule cont) have "1 / (1 + real (max M N)) \ 1 / (real M)" by (simp add: \0 < M\ frac_le) then have "\ (b \ a+1 / (1 + real (max M N)))" using M \a < b\ \M > 0\ max.cobounded1 [of M N] by linarith then show "?\ b = 0" by simp have null_a: "{a} \ null_sets (lebesgue_on {a..b})" using \a < b\ by (simp add: null_sets_restrict_space) have "LINT x|lebesgue_on {a..b}. ?g x = 0" by (force intro: integral_eq_zero_AE AE_I' [OF null_a]) then have "l2norm {a..b} ?\ < sqrt (e^2)" unfolding l2norm_def l2product_def power2_eq_square [symmetric] apply (intro real_sqrt_less_mono) using "*" by linarith then show "l2norm {a..b} ?\ < e" using \e > 0\ by auto qed auto qed lemma square_integrable_approximate_continuous_ends: assumes f: "f square_integrable {a..b}" and "a < b" "0 < e" obtains g where "continuous_on {a..b} g" "g b = g a" "g square_integrable {a..b}" "l2norm {a..b} (\x. f x - g x) < e" proof - obtain g where contg: "continuous_on UNIV g" and "g square_integrable {a..b}" and lg: "l2norm {a..b} (\x. f x - g x) < e/2" using f \e > 0\ square_integrable_approximate_continuous by (metis (full_types) box_real(2) half_gt_zero_iff lmeasurable_cbox) obtain h where conth: "continuous_on {a..b} h" and "h a = g b - g a" "h b = 0" and lh: "l2norm {a..b} h < e/2" using integral_tweak_ends \e > 0\ by (metis \a < b\ zero_less_divide_iff zero_less_numeral) have "h square_integrable {a..b}" using \continuous_on {a..b} h\ continuous_imp_square_integrable by blast show thesis proof show "continuous_on {a..b} (\x. g x + h x)" by (blast intro: continuous_on_subset [OF contg] conth continuous_intros) then show "(\x. g x + h x) square_integrable {a..b}" using continuous_imp_square_integrable by blast show "g b + h b = g a + h a" by (simp add: \h a = g b - g a\ \h b = 0\) have "l2norm {a..b} (\x. (f x - g x) + - h x) < e" proof (rule le_less_trans [OF l2norm_triangle [of "\x. f x - g x" "{a..b}" "\x. - (h x)"]]) show "(\x. f x - g x) square_integrable {a..b}" using \g square_integrable {a..b}\ f square_integrable_diff by blast show "(\x. - h x) square_integrable {a..b}" by (simp add: \h square_integrable {a..b}\) show "l2norm {a..b} (\x. f x - g x) + l2norm {a..b} (\x. - h x) < e" using \h square_integrable {a..b}\ l2norm_neg lg lh by auto qed then show "l2norm {a..b} (\x. f x - (g x + h x)) < e" by (simp add: field_simps) qed qed subsection\Hence the main approximation result\ lemma Weierstrass_l2_trig_polynomial: assumes f: "f square_integrable {-pi..pi}" and "0 < e" obtains n a b where "l2norm {-pi..pi} (\x. f x - (\k\n. a k * sin(real k * x) + b k * cos(real k * x))) < e" proof - let ?\ = "\n a b x. \k\n. a k * sin(real k * x) + b k * cos(real k * x)" obtain g where contg: "continuous_on {-pi..pi} g" and geq: "g(-pi) = g pi" and g: "g square_integrable {-pi..pi}" and norm_fg: "l2norm {-pi..pi} (\x. f x - g x) < e/2" using \e > 0\ by (auto intro: square_integrable_approximate_continuous_ends [OF f, of "e/2"]) then obtain n a b where g_phi_less: "\x. x \ {-pi..pi} \ \g x - (?\ n a b x)\ < e/6" using \e > 0\ Weierstrass_trig_polynomial [OF contg geq, of "e/6"] by (meson zero_less_divide_iff zero_less_numeral) show thesis proof have si: "(?\ n u v) square_integrable {-pi..pi}" for n::nat and u v proof (intro square_integrable_sum continuous_imp_square_integrable) show "continuous_on {-pi..pi} (\x. u k * sin (real k * x) + v k * cos (real k * x))" if "k \ {..n}" for k using that by (intro continuous_intros) qed auto have "l2norm {-pi..pi} (\x. f x - (?\ n a b x)) = l2norm {-pi..pi} (\x. (f x - g x) + (g x - (?\ n a b x)))" by simp also have "\ \ l2norm {-pi..pi} (\x. f x - g x) + l2norm {-pi..pi} (\x. g x - (?\ n a b x))" proof (rule l2norm_triangle) show "(\x. f x - g x) square_integrable {-pi..pi}" using f g square_integrable_diff by blast show "(\x. g x - (?\ n a b x)) square_integrable {-pi..pi}" using g si square_integrable_diff by blast qed also have "\ < e" proof - have g_phi: "(\x. g x - (?\ n a b x)) square_integrable {-pi..pi}" using g si square_integrable_diff by blast have "l2norm {-pi..pi} (\x. g x - (?\ n a b x)) \ e/2" unfolding l2norm_def l2product_def power2_eq_square [symmetric] proof (rule real_le_lsqrt) - show "0 \ LINT x|lebesgue_on {-pi..pi}. (g x - (?\ n a b x))\<^sup>2" - by (rule integral_nonneg_AE) auto have "LINT x|lebesgue_on {-pi..pi}. (g x - (?\ n a b x))\<^sup>2 \ LINT x|lebesgue_on {-pi..pi}. (e / 6) ^ 2" proof (rule integral_mono) show "integrable (lebesgue_on {-pi..pi}) (\x. (g x - (?\ n a b x))\<^sup>2)" using g_phi square_integrable_def by auto show "integrable (lebesgue_on {-pi..pi}) (\x. (e / 6)\<^sup>2)" by (intro continuous_intros continuous_imp_integrable_real) show "(g x - (?\ n a b x))\<^sup>2 \ (e / 6)\<^sup>2" if "x \ space (lebesgue_on {-pi..pi})" for x using \e > 0\ g_phi_less that apply (simp add: abs_le_square_iff [symmetric]) using less_eq_real_def by blast qed also have "\ \ (e / 2)\<^sup>2" using \e > 0\ pi_less_4 by (auto simp: power2_eq_square measure_restrict_space) finally show "LINT x|lebesgue_on {-pi..pi}. (g x - (?\ n a b x))\<^sup>2 \ (e / 2)\<^sup>2" . qed (use \e > 0\ in auto) with norm_fg show ?thesis by auto qed finally show "l2norm {-pi..pi} (\x. f x - (?\ n a b x)) < e" . qed qed proposition Weierstrass_l2_trigonometric_set: assumes f: "f square_integrable {-pi..pi}" and "0 < e" obtains n a where "l2norm {-pi..pi} (\x. f x - (\k\n. a k * trigonometric_set k x)) < e" proof - obtain n a b where lee: "l2norm {-pi..pi} (\x. f x - (\k\n. a k * sin(real k * x) + b k * cos(real k * x))) < e" using Weierstrass_l2_trig_polynomial [OF assms] . let ?a = "\k. if k = 0 then sqrt(2 * pi) * b 0 else if even k then sqrt pi * b(k div 2) else if k \ 2 * n then sqrt pi * a((Suc k) div 2) else 0" show thesis proof have [simp]: "Suc (i * 2) \ n * 2 \ i {..k\n. b k * cos (real k * x)) = (\i\n. if i = 0 then b 0 else b i * cos (real i * x))" for x by (rule sum.cong) auto moreover have "(\k\n. a k * sin (real k * x)) = (\i\n. (if Suc (2 * i) \ 2 * n then sqrt pi * a (Suc i) * sin ((1 + real i) * x) else 0) / sqrt pi)" (is "?lhs = ?rhs") for x proof (cases "n=0") case False then obtain n' where n': "n = Suc n'" using not0_implies_Suc by blast have "?lhs = (\k = Suc 0..n. a k * sin (real k * x))" by (simp add: atMost_atLeast0 sum_shift_lb_Suc0_0) also have "\ = (\ij\n'. a (Suc j) * sin (real (Suc j) * x)) = (\i = ?rhs" by (simp add: field_simps if_distrib [of "\x. x/_"] sum.inter_restrict [where B = "{..k\n. a k * sin(real k * x) + b k * cos(real k * x)) = (\k \ Suc(2*n). ?a k * trigonometric_set k x)" for x unfolding sum.in_pairs_0 trigonometric_set_def by (simp add: sum.distrib if_distrib [of "\x. x*_"] cong: if_cong) with lee show "l2norm {-pi..pi} (\x. f x - (\k \ Suc(2*n). ?a k * trigonometric_set k x)) < e" by auto qed qed subsection\Convergence wrt the L2 norm of trigonometric Fourier series\ definition Fourier_coefficient where "Fourier_coefficient \ orthonormal_coeff {-pi..pi} trigonometric_set" lemma Fourier_series_l2: assumes "f square_integrable {-pi..pi}" shows "(\n. l2norm {-pi..pi} (\x. f x - (\i\n. Fourier_coefficient f i * trigonometric_set i x))) \ 0" proof (clarsimp simp add: lim_sequentially dist_real_def Fourier_coefficient_def) let ?h = "\n x. (\i\n. orthonormal_coeff {-pi..pi} trigonometric_set f i * trigonometric_set i x)" show "\N. \n\N. \l2norm {-pi..pi} (\x. f x - ?h n x)\ < e" if "0 < e" for e proof - obtain N a where lte: "l2norm {-pi..pi} (\x. f x - (\k\N. a k * trigonometric_set k x)) < e" using Weierstrass_l2_trigonometric_set by (meson \0 < e\ assms) show ?thesis proof (intro exI allI impI) show "\l2norm {-pi..pi} (\x. f x - ?h m x)\ < e" if "N \ m" for m proof - have "\l2norm {-pi..pi} (\x. f x - ?h m x)\ = l2norm {-pi..pi} (\x. f x - ?h m x)" proof (rule abs_of_nonneg) show "0 \ l2norm {-pi..pi} (\x. f x - ?h m x)" apply (intro l2norm_pos_le square_integrable_diff square_integrable_sum square_integrable_lmult square_integrable_trigonometric_set assms, auto) done qed also have "\ \ l2norm {-pi..pi} (\x. f x - (\k\N. a k * trigonometric_set k x))" proof - have "(\i\m. (if i \ N then a i else 0) * trigonometric_set i x) = (\i\N. a i * trigonometric_set i x)" for x using sum.inter_restrict [where A = "{..m}" and B = "{..N}", symmetric] that by (force simp: if_distrib [of "\x. x * _"] min_absorb2 cong: if_cong) moreover have "l2norm {-pi..pi} (\x. f x - ?h m x) \ l2norm {-pi..pi} (\x. f x - (\i\m. (if i \ N then a i else 0) * trigonometric_set i x))" using orthonormal_optimal_partial_sum [OF orthonormal_system_trigonometric_set square_integrable_trigonometric_set assms] by simp ultimately show ?thesis by simp qed also have "\ < e" by (rule lte) finally show ?thesis . qed qed qed qed subsection\Fourier coefficients go to 0 (weak form of Riemann-Lebesgue)\ lemma trigonometric_set_mul_absolutely_integrable: assumes "f absolutely_integrable_on {-pi..pi}" shows "(\x. trigonometric_set n x * f x) absolutely_integrable_on {-pi..pi}" proof (rule absolutely_integrable_bounded_measurable_product_real) show "trigonometric_set n \ borel_measurable (lebesgue_on {-pi..pi})" using square_integrable_def square_integrable_trigonometric_set by blast show "bounded (trigonometric_set n ` {-pi..pi})" unfolding bounded_iff using pi_gt3 sqrt_pi_ge1 by (rule_tac x=1 in exI) (auto simp: trigonometric_set_def dist_real_def intro: order_trans [OF abs_sin_le_one] order_trans [OF abs_cos_le_one]) qed (auto simp: assms) lemma trigonometric_set_mul_integrable: "f absolutely_integrable_on {-pi..pi} \ integrable (lebesgue_on {-pi..pi}) (\x. trigonometric_set n x * f x)" using trigonometric_set_mul_absolutely_integrable by (simp add: integrable_restrict_space set_integrable_def) lemma trigonometric_set_integrable [simp]: "integrable (lebesgue_on {-pi..pi}) (trigonometric_set n)" using trigonometric_set_mul_integrable [where f = id] by simp (metis absolutely_integrable_imp_integrable fmeasurableD interval_cbox lmeasurable_cbox square_integrable_imp_absolutely_integrable square_integrable_trigonometric_set) lemma absolutely_integrable_sin_product: assumes "f absolutely_integrable_on {-pi..pi}" shows "(\x. sin(k * x) * f x) absolutely_integrable_on {-pi..pi}" proof (rule absolutely_integrable_bounded_measurable_product_real) show "(\x. sin (k * x)) \ borel_measurable (lebesgue_on {-pi..pi})" by (metis borel_measurable_integrable integrable_sin_cx mult_commute_abs) show "bounded ((\x. sin (k * x)) ` {-pi..pi})" by (metis (mono_tags, lifting) abs_sin_le_one bounded_iff imageE real_norm_def) qed (auto simp: assms) lemma absolutely_integrable_cos_product: assumes "f absolutely_integrable_on {-pi..pi}" shows "(\x. cos(k * x) * f x) absolutely_integrable_on {-pi..pi}" proof (rule absolutely_integrable_bounded_measurable_product_real) show "(\x. cos (k * x)) \ borel_measurable (lebesgue_on {-pi..pi})" by (metis borel_measurable_integrable integrable_cos_cx mult_commute_abs) show "bounded ((\x. cos (k * x)) ` {-pi..pi})" by (metis (mono_tags, lifting) abs_cos_le_one bounded_iff imageE real_norm_def) qed (auto simp: assms) lemma assumes "f absolutely_integrable_on {-pi..pi}" shows Fourier_products_integrable_cos: "integrable (lebesgue_on {-pi..pi}) (\x. cos(k * x) * f x)" and Fourier_products_integrable_sin: "integrable (lebesgue_on {-pi..pi}) (\x. sin(k * x) * f x)" using absolutely_integrable_cos_product absolutely_integrable_sin_product assms by (auto simp: integrable_restrict_space set_integrable_def) lemma Riemann_lebesgue_square_integrable: assumes "orthonormal_system S w" "\i. w i square_integrable S" "f square_integrable S" shows "orthonormal_coeff S w f \ 0" using Fourier_series_square_summable [OF assms, of UNIV] summable_LIMSEQ_zero by force proposition Riemann_lebesgue: assumes "f absolutely_integrable_on {-pi..pi}" shows "Fourier_coefficient f \ 0" unfolding lim_sequentially proof (intro allI impI) fix e::real assume "e > 0" then obtain g where "continuous_on UNIV g" and gabs: "g absolutely_integrable_on {-pi..pi}" and fg_e2: "integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. \f x - g x\) < e/2" using absolutely_integrable_approximate_continuous [OF assms, of "e/2"] by (metis (full_types) box_real(2) half_gt_zero_iff lmeasurable_cbox) have "g square_integrable {-pi..pi}" using \continuous_on UNIV g\ continuous_imp_square_integrable continuous_on_subset by blast then have "orthonormal_coeff {-pi..pi} trigonometric_set g \ 0" using Riemann_lebesgue_square_integrable orthonormal_system_trigonometric_set square_integrable_trigonometric_set by blast with \e > 0\ obtain N where N: "\n. n \ N \ \orthonormal_coeff {-pi..pi} trigonometric_set g n\ < e/2" unfolding lim_sequentially by (metis half_gt_zero_iff norm_conv_dist real_norm_def) have "\Fourier_coefficient f n\ < e" if "n \ N" for n proof - have "\LINT x|lebesgue_on {-pi..pi}. trigonometric_set n x * g x\ < e/2" using N [OF \n \ N\] by (auto simp: orthonormal_coeff_def l2product_def) have "integrable (lebesgue_on {-pi..pi}) (\x. trigonometric_set n x * g x)" using gabs trigonometric_set_mul_integrable by blast moreover have "integrable (lebesgue_on {-pi..pi}) (\x. trigonometric_set n x * f x)" using assms trigonometric_set_mul_integrable by blast ultimately have "\(LINT x|lebesgue_on {-pi..pi}. trigonometric_set n x * g x) - (LINT x|lebesgue_on {-pi..pi}. trigonometric_set n x * f x)\ = \(LINT x|lebesgue_on {-pi..pi}. trigonometric_set n x * (g x - f x))\" by (simp add: algebra_simps flip: Bochner_Integration.integral_diff) also have "\ \ LINT x|lebesgue_on {-pi..pi}. \f x - g x\" proof (rule integral_abs_bound_integral) show "integrable (lebesgue_on {-pi..pi}) (\x. trigonometric_set n x * (g x - f x))" by (simp add: gabs assms trigonometric_set_mul_integrable) have "(\x. f x - g x) absolutely_integrable_on {-pi..pi}" using gabs assms by blast then show "integrable (lebesgue_on {-pi..pi}) (\x. \f x - g x\)" by (simp add: absolutely_integrable_imp_integrable) fix x assume "x \ space (lebesgue_on {-pi..pi})" then have "-pi \ x" "x \ pi" by auto have "\trigonometric_set n x\ \ 1" using pi_ge_two apply (simp add: trigonometric_set_def) using sqrt_pi_ge1 abs_sin_le_one order_trans abs_cos_le_one by metis then show "\trigonometric_set n x * (g x - f x)\ \ \f x - g x\" using abs_ge_zero mult_right_mono by (fastforce simp add: abs_mult abs_minus_commute) qed finally have "\(LINT x|lebesgue_on {-pi..pi}. trigonometric_set n x * g x) - (LINT x|lebesgue_on {-pi..pi}. trigonometric_set n x * f x)\ \ LINT x|lebesgue_on {-pi..pi}. \f x - g x\" . then show ?thesis using N [OF \n \ N\] fg_e2 unfolding Fourier_coefficient_def orthonormal_coeff_def l2product_def by linarith qed then show "\no. \n\no. dist (Fourier_coefficient f n) 0 < e" by auto qed lemma Riemann_lebesgue_sin: assumes "f absolutely_integrable_on {-pi..pi}" shows "(\n. integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. sin(real n * x) * f x)) \ 0" unfolding lim_sequentially proof (intro allI impI) fix e::real assume "e > 0" then obtain N where N: "\n. n \ N \ \Fourier_coefficient f n\ < e/4" using Riemann_lebesgue [OF assms] unfolding lim_sequentially by (metis norm_conv_dist real_norm_def zero_less_divide_iff zero_less_numeral) have "\LINT x|lebesgue_on {-pi..pi}. sin (real n * x) * f x\ < e" if "n > N" for n using that proof (induction n) case (Suc n) have "\Fourier_coefficient f(Suc (2 * n))\ < e/4" using N Suc.prems by auto then have "\LINT x|lebesgue_on {-pi..pi}. sin ((1 + real n) * x) * f x\ < sqrt pi * e / 4" by (simp add: Fourier_coefficient_def orthonormal_coeff_def trigonometric_set_def l2product_def field_simps) also have "\ \ e" using \0 < e\ pi_less_4 real_sqrt_less_mono by (fastforce simp add: field_simps) finally show ?case by simp qed auto then show "\no. \n\no. dist (LINT x|lebesgue_on {-pi..pi}. sin (real n * x) * f x) 0 < e" by (metis (full_types) le_neq_implies_less less_add_same_cancel2 less_trans norm_conv_dist real_norm_def zero_less_one) qed lemma Riemann_lebesgue_cos: assumes "f absolutely_integrable_on {-pi..pi}" shows "(\n. integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. cos(real n * x) * f x)) \ 0" unfolding lim_sequentially proof (intro allI impI) fix e::real assume "e > 0" then obtain N where N: "\n. n \ N \ \Fourier_coefficient f n\ < e/4" using Riemann_lebesgue [OF assms] unfolding lim_sequentially by (metis norm_conv_dist real_norm_def zero_less_divide_iff zero_less_numeral) have "\LINT x|lebesgue_on {-pi..pi}. cos (real n * x) * f x\ < e" if "n > N" for n using that proof (induction n) case (Suc n) have eq: "(x * 2 + x * (real n * 2)) / 2 = x + x * (real n)" for x by simp have "\Fourier_coefficient f(2*n + 2)\ < e/4" using N Suc.prems by auto then have "\LINT x|lebesgue_on {-pi..pi}. f x * cos (x + x * (real n))\ < sqrt pi * e / 4" by (simp add: Fourier_coefficient_def orthonormal_coeff_def trigonometric_set_def l2product_def field_simps eq) also have "\ \ e" using \0 < e\ pi_less_4 real_sqrt_less_mono by (fastforce simp add: field_simps) finally show ?case by (simp add: field_simps) qed auto then show "\no. \n\no. dist (LINT x|lebesgue_on {-pi..pi}. cos (real n * x) * f x) 0 < e" by (metis (full_types) le_neq_implies_less less_add_same_cancel2 less_trans norm_conv_dist real_norm_def zero_less_one) qed lemma Riemann_lebesgue_sin_half: assumes "f absolutely_integrable_on {-pi..pi}" shows "(\n. LINT x|lebesgue_on {-pi..pi}. sin ((real n + 1/2) * x) * f x) \ 0" proof (simp add: algebra_simps sin_add) let ?INT = "integral\<^sup>L (lebesgue_on {-pi..pi})" let ?f = "(\n. ?INT (\x. sin(n * x) * cos(1/2 * x) * f x) + ?INT (\x. cos(n * x) * sin(1/2 * x) * f x))" show "(\n. ?INT (\x. f x * (cos (x * real n) * sin (x/2)) + f x * (sin (x * real n) * cos (x/2)))) \ 0" proof (rule Lim_transform_eventually) have sin: "(\x. sin (1/2 * x) * f x) absolutely_integrable_on {-pi..pi}" by (intro absolutely_integrable_sin_product assms) have cos: "(\x. cos (1/2 * x) * f x) absolutely_integrable_on {-pi..pi}" by (intro absolutely_integrable_cos_product assms) show "\\<^sub>F n in sequentially. ?f n = ?INT (\x. f x * (cos (x * real n) * sin (x/2)) + f x * (sin (x * real n) * cos (x/2)))" unfolding mult.assoc apply (rule eventuallyI) apply (subst Bochner_Integration.integral_add [symmetric]) apply (safe intro!: cos absolutely_integrable_sin_product sin absolutely_integrable_cos_product absolutely_integrable_imp_integrable) apply (auto simp: algebra_simps) done have "?f \ 0 + 0" unfolding mult.assoc by (intro tendsto_add Riemann_lebesgue_sin Riemann_lebesgue_cos sin cos) then show "?f \ (0::real)" by simp qed qed lemma Fourier_sum_limit_pair: assumes "f absolutely_integrable_on {-pi..pi}" shows "(\n. \k\2 * n. Fourier_coefficient f k * trigonometric_set k t) \ l \ (\n. \k\n. Fourier_coefficient f k * trigonometric_set k t) \ l" (is "?lhs = ?rhs") proof assume L: ?lhs show ?rhs unfolding lim_sequentially dist_real_def proof (intro allI impI) fix e::real assume "e > 0" then obtain N1 where N1: "\n. n \ N1 \ \Fourier_coefficient f n\ < e/2" using Riemann_lebesgue [OF assms] unfolding lim_sequentially by (metis norm_conv_dist real_norm_def zero_less_divide_iff zero_less_numeral) obtain N2 where N2: "\n. n \ N2 \ \(\k\2 * n. Fourier_coefficient f k * trigonometric_set k t) - l\ < e/2" using L unfolding lim_sequentially dist_real_def by (meson \0 < e\ half_gt_zero_iff) show "\no. \n\no. \(\k\n. Fourier_coefficient f k * trigonometric_set k t) - l\ < e" proof (intro exI allI impI) fix n assume n: "N1 + 2*N2 + 1 \ n" then have "n \ N1" "n \ N2" "n div 2 \ N2" by linarith+ consider "n = 2 * (n div 2)" | "n = Suc(2 * (n div 2))" by linarith then show "\(\k\n. Fourier_coefficient f k * trigonometric_set k t) - l\ < e" proof cases case 1 show ?thesis apply (subst 1) using N2 [OF \n div 2 \ N2\] by linarith next case 2 have "\(\k\2 * (n div 2). Fourier_coefficient f k * trigonometric_set k t) - l\ < e/2" using N2 [OF \n div 2 \ N2\] by linarith moreover have "\Fourier_coefficient f(Suc (2 * (n div 2))) * trigonometric_set (Suc (2 * (n div 2))) t\ < (e/2) * 1" proof - have "\trigonometric_set (Suc (2 * (n div 2))) t\ \ 1" apply (simp add: trigonometric_set) using abs_sin_le_one sqrt_pi_ge1 by (blast intro: order_trans) moreover have "\Fourier_coefficient f(Suc (2 * (n div 2)))\ < e/2" using N1 \N1 \ n\ by auto ultimately show ?thesis unfolding abs_mult by (meson abs_ge_zero le_less_trans mult_left_mono mult_less_iff1 zero_less_one) qed ultimately show ?thesis apply (subst 2) unfolding sum.atMost_Suc by linarith qed qed qed next assume ?rhs then show ?lhs unfolding lim_sequentially by (auto simp: elim!: imp_forward ex_forward) qed subsection\Express Fourier sum in terms of the special expansion at the origin\ lemma Fourier_sum_0: "(\k \ n. Fourier_coefficient f k * trigonometric_set k 0) = (\k \ n div 2. Fourier_coefficient f(2*k) * trigonometric_set (2*k) 0)" (is "?lhs = ?rhs") proof - have "?lhs = (\k = 2 * 0.. Suc (2 * (n div 2)). Fourier_coefficient f k * trigonometric_set k 0)" proof (rule sum.mono_neutral_left) show "\i\{2 * 0..Suc (2 * (n div 2))} - {..n}. Fourier_coefficient f i * trigonometric_set i 0 = 0" proof clarsimp fix i assume "i \ Suc (2 * (n div 2))" "\ i \ n" then have "i = Suc n" "even n" by presburger+ moreover assume "trigonometric_set i 0 \ 0" ultimately show "Fourier_coefficient f i = 0" by (simp add: trigonometric_set_def) qed qed auto also have "\ = ?rhs" unfolding sum.in_pairs by (simp add: trigonometric_set atMost_atLeast0) finally show ?thesis . qed lemma Fourier_sum_0_explicit: "(\k\n. Fourier_coefficient f k * trigonometric_set k 0) = (Fourier_coefficient f 0 / sqrt 2 + (\k = 1..n div 2. Fourier_coefficient f(2*k))) / sqrt pi" (is "?lhs = ?rhs") proof - have "(\k=0..n. Fourier_coefficient f k * trigonometric_set k 0) = Fourier_coefficient f 0 * trigonometric_set 0 0 + (\k = 1..n. Fourier_coefficient f k * trigonometric_set k 0)" by (simp add: Fourier_sum_0 sum.atLeast_Suc_atMost) also have "\ = ?rhs" proof - have "Fourier_coefficient f 0 * trigonometric_set 0 0 = Fourier_coefficient f 0 / (sqrt 2 * sqrt pi)" by (simp add: real_sqrt_mult trigonometric_set(1)) moreover have "(\k = Suc 0..n. Fourier_coefficient f k * trigonometric_set k 0) = (\k = Suc 0..n div 2. Fourier_coefficient f(2*k)) / sqrt pi" proof (induction n) case (Suc n) show ?case by (simp add: Suc) (auto simp: trigonometric_set_def field_simps) qed auto ultimately show ?thesis by (simp add: add_divide_distrib) qed finally show ?thesis by (simp add: atMost_atLeast0) qed lemma Fourier_sum_0_integrals: assumes "f absolutely_integrable_on {-pi..pi}" shows "(\k\n. Fourier_coefficient f k * trigonometric_set k 0) = (integral\<^sup>L (lebesgue_on {-pi..pi}) f / 2 + (\k = Suc 0..n div 2. integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. cos(k * x) * f x))) / pi" proof - have "integral\<^sup>L (lebesgue_on {-pi..pi}) f / (sqrt (2 * pi) * sqrt 2 * sqrt pi) = integral\<^sup>L (lebesgue_on {-pi..pi}) f / (2 * pi)" by (simp add: algebra_simps real_sqrt_mult) moreover have "(\k = Suc 0..n div 2. LINT x|lebesgue_on {-pi..pi}. trigonometric_set (2*k) x * f x) / sqrt pi = (\k = Suc 0..n div 2. LINT x|lebesgue_on {-pi..pi}. cos (k * x) * f x) / pi" by (simp add: trigonometric_set_def sum_divide_distrib) ultimately show ?thesis unfolding Fourier_sum_0_explicit by (simp add: Fourier_coefficient_def orthonormal_coeff_def l2product_def trigonometric_set add_divide_distrib) qed lemma Fourier_sum_0_integral: assumes "f absolutely_integrable_on {-pi..pi}" shows "(\k\n. Fourier_coefficient f k * trigonometric_set k 0) = integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. (1/2 + (\k = Suc 0..n div 2. cos(k * x))) * f x) / pi" proof - note * = Fourier_products_integrable_cos [OF assms] have "integrable (lebesgue_on {-pi..pi}) (\x. \n = Suc 0..n div 2. f x * cos (x * real n))" using * by (force simp: mult_ac) moreover have "integral\<^sup>L (lebesgue_on {-pi..pi}) f / 2 + (\k = Suc 0..n div 2. LINT x|lebesgue_on {-pi..pi}. f x * cos (x * real k)) = (LINT x|lebesgue_on {-pi..pi}. f x / 2) + (LINT x|lebesgue_on {-pi..pi}. (\n = Suc 0..n div 2. f x * cos (x * real n)))" proof (subst Bochner_Integration.integral_sum) show "integrable (lebesgue_on {-pi..pi}) (\x. f x * cos (x * real i))" if "i \ {Suc 0..n div 2}" for i using that * [of i] by (auto simp: Bochner_Integration.integral_sum mult_ac) qed auto ultimately show ?thesis using Fourier_products_integrable_cos [OF assms] absolutely_integrable_imp_integrable [OF assms] by (simp add: Fourier_sum_0_integrals sum_distrib_left assms algebra_simps) qed subsection\How Fourier coefficients behave under addition etc\ lemma Fourier_coefficient_add: assumes "f absolutely_integrable_on {-pi..pi}" "g absolutely_integrable_on {-pi..pi}" shows "Fourier_coefficient (\x. f x + g x) i = Fourier_coefficient f i + Fourier_coefficient g i" using assms trigonometric_set_mul_integrable by (simp add: Fourier_coefficient_def orthonormal_coeff_def l2product_def distrib_left) lemma Fourier_coefficient_minus: assumes "f absolutely_integrable_on {-pi..pi}" shows "Fourier_coefficient (\x. - f x) i = - Fourier_coefficient f i" using assms trigonometric_set_mul_integrable by (simp add: Fourier_coefficient_def orthonormal_coeff_def l2product_def) lemma Fourier_coefficient_diff: assumes f: "f absolutely_integrable_on {-pi..pi}" and g: "g absolutely_integrable_on {-pi..pi}" shows "Fourier_coefficient (\x. f x - g x) i = Fourier_coefficient f i - Fourier_coefficient g i" proof - have mg: "(\x. - g x) absolutely_integrable_on {-pi..pi}" using g by (simp add: absolutely_integrable_measurable_real) show ?thesis using Fourier_coefficient_add [OF f mg] Fourier_coefficient_minus [OF g] by simp qed lemma Fourier_coefficient_const: "Fourier_coefficient (\x. c) i = (if i = 0 then c * sqrt(2 * pi) else 0)" by (auto simp: Fourier_coefficient_def orthonormal_coeff_def l2product_def trigonometric_set_def divide_simps measure_restrict_space) lemma Fourier_offset_term: fixes f :: "real \ real" assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" shows "Fourier_coefficient (\x. f(x+t)) (2 * n + 2) * trigonometric_set (2 * n + 2) 0 = Fourier_coefficient f(2 * n+1) * trigonometric_set (2 * n+1) t + Fourier_coefficient f(2 * n + 2) * trigonometric_set (2 * n + 2) t" proof - have eq: "(2 + 2 * real n) * x / 2 = (1 + real n) * x" for x by (simp add: divide_simps) have 1: "integrable (lebesgue_on {-pi..pi}) (\x. f x * (cos (x + x * n) * cos (t + t * n)))" using Fourier_products_integrable_cos [of f "Suc n"] f by (simp add: algebra_simps) (simp add: mult.assoc [symmetric]) have 2: "integrable (lebesgue_on {-pi..pi}) (\x. f x * (sin (x + x * n) * sin (t + t * n)))" using Fourier_products_integrable_sin [of f "Suc n"] f by (simp add: algebra_simps) (simp add: mult.assoc [symmetric]) have "has_bochner_integral (lebesgue_on {-pi..pi}) (\x. cos (real (Suc n) * (x + t - t)) * f(x + t)) (LINT x|lebesgue_on {-pi..pi}. cos (real (Suc n) * (x - t)) * f x)" proof (rule has_integral_periodic_offset) have "(\x. cos (real (Suc n) * (x - t)) * f x) absolutely_integrable_on {-pi..pi}" proof (rule absolutely_integrable_bounded_measurable_product_real) show "(\x. cos (real (Suc n) * (x - t))) \ borel_measurable (lebesgue_on {-pi..pi})" by (intro continuous_imp_measurable_on_sets_lebesgue continuous_intros) auto show "bounded ((\x. cos (real (Suc n) * (x - t))) ` {-pi..pi})" by (rule boundedI [where B=1]) auto qed (use f in auto) then show "has_bochner_integral (lebesgue_on {-pi..pi}) (\x. cos ((Suc n) * (x - t)) * f x) (LINT x|lebesgue_on {-pi..pi}. cos (real (Suc n) * (x - t)) * f x)" by (simp add: has_bochner_integral_integrable integrable_restrict_space set_integrable_def) next fix x show "cos (real (Suc n) * (x + (pi - - pi) - t)) * f(x + (pi - - pi)) = cos (real (Suc n) * (x - t)) * f x" using periodic cos.plus_of_nat [of "(Suc n) * (x - t)" "Suc n"] by (simp add: algebra_simps) qed then have "has_bochner_integral (lebesgue_on {-pi..pi}) (\x. cos (real (Suc n) * x) * f(x + t)) (LINT x|lebesgue_on {-pi..pi}. cos (real (Suc n) * (x - t)) * f x)" by simp then have "LINT x|lebesgue_on {-pi..pi}. cos ((Suc n) * x) * f(x+t) = LINT x|lebesgue_on {-pi..pi}. cos ((Suc n) * (x - t)) * f x" using has_bochner_integral_integral_eq by blast also have "\ = LINT x|lebesgue_on {-pi..pi}. cos ((Suc n) * x - ((Suc n) * t)) * f x" by (simp add: algebra_simps) also have "\ = cos ((Suc n) * t) * (LINT x|lebesgue_on {-pi..pi}. cos ((Suc n) * x) * f x) + sin ((Suc n) * t) * (LINT x|lebesgue_on {-pi..pi}. sin ((Suc n) * x) * f x)" by (simp add: cos_diff algebra_simps 1 2 flip: integral_mult_left_zero integral_mult_right_zero) finally have "LINT x|lebesgue_on {-pi..pi}. cos ((Suc n) * x) * f(x+t) = cos ((Suc n) * t) * (LINT x|lebesgue_on {-pi..pi}. cos ((Suc n) * x) * f x) + sin ((Suc n) * t) * (LINT x|lebesgue_on {-pi..pi}. sin ((Suc n) * x) * f x)" . then show ?thesis unfolding Fourier_coefficient_def orthonormal_coeff_def trigonometric_set_def by (simp add: eq) (simp add: divide_simps algebra_simps l2product_def) qed lemma Fourier_sum_offset: fixes f :: "real \ real" assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" shows "(\k\2*n. Fourier_coefficient f k * trigonometric_set k t) = (\k\2*n. Fourier_coefficient (\x. f(x+t)) k * trigonometric_set k 0)" (is "?lhs = ?rhs") proof - have "?lhs = Fourier_coefficient f 0 * trigonometric_set 0 t + (\k = Suc 0..2*n. Fourier_coefficient f k * trigonometric_set k t)" by (simp add: atMost_atLeast0 sum.atLeast_Suc_atMost) moreover have "(\k = Suc 0..2*n. Fourier_coefficient f k * trigonometric_set k t) = (\k = Suc 0..2*n. Fourier_coefficient (\x. f(x+t)) k * trigonometric_set k 0)" proof (cases n) case (Suc n') have "(\k = Suc 0..2 * n. Fourier_coefficient f k * trigonometric_set k t) = (\k = Suc 0.. Suc (Suc (2 * n')). Fourier_coefficient f k * trigonometric_set k t)" by (simp add: Suc) also have "\ = (\k \ Suc (2 * n'). Fourier_coefficient f(Suc k) * trigonometric_set (Suc k) t)" unfolding atMost_atLeast0 using sum.shift_bounds_cl_Suc_ivl by blast also have "\ = (\k \ Suc (2 * n'). Fourier_coefficient (\x. f(x+t)) (Suc k) * trigonometric_set (Suc k) 0)" unfolding sum.in_pairs_0 proof (rule sum.cong [OF refl]) show "Fourier_coefficient f(Suc (2 * k)) * trigonometric_set (Suc (2 * k)) t + Fourier_coefficient f(Suc (Suc (2 * k))) * trigonometric_set (Suc (Suc (2 * k))) t = Fourier_coefficient (\x. f(x + t)) (Suc (2 * k)) * trigonometric_set (Suc (2 * k)) 0 + Fourier_coefficient (\x. f(x + t)) (Suc (Suc (2 * k))) * trigonometric_set (Suc (Suc (2 * k))) 0" if "k \ {..n'}" for k using that Fourier_offset_term [OF assms, of t ] by (auto simp: trigonometric_set_def) qed also have "\ = (\k = Suc 0..Suc (Suc (2 * n')). Fourier_coefficient (\x. f(x+t)) k * trigonometric_set k 0)" by (simp only: atMost_atLeast0 sum.shift_bounds_cl_Suc_ivl) also have "\ = (\k = Suc 0..2*n. Fourier_coefficient (\x. f(x+t)) k * trigonometric_set k 0)" by (simp add: Suc) finally show ?thesis . qed auto moreover have "?rhs = Fourier_coefficient (\x. f(x+t)) 0 * trigonometric_set 0 0 + (\k = Suc 0..2*n. Fourier_coefficient (\x. f(x+t)) k * trigonometric_set k 0)" by (simp add: atMost_atLeast0 sum.atLeast_Suc_atMost) moreover have "Fourier_coefficient f 0 * trigonometric_set 0 t = Fourier_coefficient (\x. f(x+t)) 0 * trigonometric_set 0 0" by (simp add: Fourier_coefficient_def orthonormal_coeff_def trigonometric_set_def l2product_def integral_periodic_offset periodic) ultimately show ?thesis by simp qed lemma Fourier_sum_offset_unpaired: fixes f :: "real \ real" assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" shows "(\k\2*n. Fourier_coefficient f k * trigonometric_set k t) = (\k\n. Fourier_coefficient (\x. f(x+t)) (2*k) * trigonometric_set (2*k) 0)" (is "?lhs = ?rhs") proof - have "?lhs = (\k\n. Fourier_coefficient (\x. f(x+t)) (2*k) * trigonometric_set (2*k) 0 + Fourier_coefficient (\x. f(x+t)) (Suc (2*k)) * trigonometric_set (Suc (2*k)) 0)" apply (simp add: assms Fourier_sum_offset) apply (subst sum.in_pairs_0 [symmetric]) by (simp add: trigonometric_set) also have "\ = ?rhs" by (auto simp: trigonometric_set) finally show ?thesis . qed subsection\Express partial sums using Dirichlet kernel\ definition Dirichlet_kernel where "Dirichlet_kernel \ \n x. if x = 0 then real n + 1/2 else sin((real n + 1/2) * x) / (2 * sin(x/2))" lemma Dirichlet_kernel_0 [simp]: "\x\ < 2 * pi \ Dirichlet_kernel 0 x = 1/2" by (auto simp: Dirichlet_kernel_def sin_zero_iff) lemma Dirichlet_kernel_minus [simp]: "Dirichlet_kernel n (-x) = Dirichlet_kernel n x" by (auto simp: Dirichlet_kernel_def) lemma Dirichlet_kernel_continuous_strong: "continuous_on {-(2 * pi)<..<2 * pi} (Dirichlet_kernel n)" proof - have "isCont (Dirichlet_kernel n) z" if "-(2 * pi) < z" "z < 2 * pi" for z proof (cases "z=0") case True have *: "(\x. sin ((n + 1/2) * x) / (2 * sin (x/2))) \0 \ real n + 1/2" by real_asymp show ?thesis unfolding Dirichlet_kernel_def continuous_at True apply (rule Lim_transform_eventually [where f = "\x. sin((n + 1/2) * x) / (2 * sin(x/2))"]) apply (auto simp: * eventually_at_filter) done next case False have "continuous (at z) (\x. sin((real n + 1/2) * x) / (2 * sin(x/2)))" using that False by (intro continuous_intros) (auto simp: sin_zero_iff) moreover have "\\<^sub>F x in nhds z. x \ 0" using False t1_space_nhds by blast ultimately show ?thesis using False by (force simp: Dirichlet_kernel_def continuous_at eventually_at_filter elim: Lim_transform_eventually) qed then show ?thesis by (simp add: continuous_on_eq_continuous_at) qed lemma Dirichlet_kernel_continuous: "continuous_on {-pi..pi} (Dirichlet_kernel n)" apply (rule continuous_on_subset [OF Dirichlet_kernel_continuous_strong], clarsimp) using pi_gt_zero by linarith lemma absolutely_integrable_mult_Dirichlet_kernel: assumes "f absolutely_integrable_on {-pi..pi}" shows "(\x. Dirichlet_kernel n x * f x) absolutely_integrable_on {-pi..pi}" proof (rule absolutely_integrable_bounded_measurable_product_real) show "Dirichlet_kernel n \ borel_measurable (lebesgue_on {-pi..pi})" by (simp add: Dirichlet_kernel_continuous continuous_imp_measurable_on_sets_lebesgue) have "compact (Dirichlet_kernel n ` {-pi..pi})" by (auto simp: compact_continuous_image [OF Dirichlet_kernel_continuous]) then show "bounded (Dirichlet_kernel n ` {-pi..pi})" using compact_imp_bounded by blast qed (use assms in auto) lemma cosine_sum_lemma: "(1/2 + (\k = Suc 0..n. cos(real k * x))) * sin(x/2) = sin((real n + 1/2) * x) / 2" proof - consider "n=0" | "n \ 1" by linarith then show ?thesis proof cases case 2 then have "(\k = Suc 0..n. (sin (real k * x + x/2) * 2 - sin (real k * x - x/2) * 2) / 2) = sin (real n * x + x/2) - sin (x/2)" proof (induction n) case (Suc n) show ?case proof (cases "n=0") case False with Suc show ?thesis by (simp add: divide_simps algebra_simps) qed auto qed auto then show ?thesis by (simp add: distrib_right sum_distrib_right cos_times_sin) qed auto qed lemma Dirichlet_kernel_cosine_sum: assumes "\x\ < 2 * pi" shows "Dirichlet_kernel n x = 1/2 + (\k = Suc 0..n. cos(real k * x))" proof - have "sin ((real n + 1/2) * x) / (2 * sin (x/2)) = 1/2 + (\k = Suc 0..n. cos (real k * x))" if "x \ 0" proof - have "sin(x/2) \ 0" using assms that by (auto simp: sin_zero_iff) then show ?thesis using cosine_sum_lemma [of x n] by (simp add: divide_simps) qed then show ?thesis by (auto simp: Dirichlet_kernel_def) qed lemma integrable_Dirichlet_kernel: "integrable (lebesgue_on {-pi..pi}) (Dirichlet_kernel n)" using Dirichlet_kernel_continuous continuous_imp_integrable_real by blast lemma integral_Dirichlet_kernel [simp]: "integral\<^sup>L (lebesgue_on {-pi..pi}) (Dirichlet_kernel n) = pi" proof - have "integral\<^sup>L (lebesgue_on {-pi..pi}) (Dirichlet_kernel n) = LINT x|lebesgue_on {-pi..pi}. 1/2 + (\k = Suc 0..n. cos (k * x))" using pi_ge_two by (force intro: Bochner_Integration.integral_cong [OF refl Dirichlet_kernel_cosine_sum]) also have "\ = (LINT x|lebesgue_on {-pi..pi}. 1/2) + (LINT x|lebesgue_on {-pi..pi}. (\k = Suc 0..n. cos (k * x)))" proof (rule Bochner_Integration.integral_add) show "integrable (lebesgue_on {-pi..pi}) (\x. \k = Suc 0..n. cos (real k * x))" by (rule Bochner_Integration.integrable_sum) (metis integrable_cos_cx mult_commute_abs) qed auto also have "\ = pi" proof - have "\i. \Suc 0 \ i; i \ n\ \ integrable (lebesgue_on {-pi..pi}) (\x. cos (real i * x))" by (metis integrable_cos_cx mult_commute_abs) then have "LINT x|lebesgue_on {-pi..pi}. (\k = Suc 0..n. cos (real k * x)) = 0" by (simp add: Bochner_Integration.integral_sum) then show ?thesis by (auto simp: measure_restrict_space) qed finally show ?thesis . qed lemma integral_Dirichlet_kernel_half [simp]: "integral\<^sup>L (lebesgue_on {0..pi}) (Dirichlet_kernel n) = pi/2" proof - have "integral\<^sup>L (lebesgue_on {- pi..0}) (Dirichlet_kernel n) + integral\<^sup>L (lebesgue_on {0..pi}) (Dirichlet_kernel n) = pi" using integral_combine [OF integrable_Dirichlet_kernel] integral_Dirichlet_kernel by simp moreover have "integral\<^sup>L (lebesgue_on {- pi..0}) (Dirichlet_kernel n) = integral\<^sup>L (lebesgue_on {0..pi}) (Dirichlet_kernel n)" using integral_reflect_real [of pi 0 "Dirichlet_kernel n"] by simp ultimately show ?thesis by simp qed lemma Fourier_sum_offset_Dirichlet_kernel: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" shows "(\k\2*n. Fourier_coefficient f k * trigonometric_set k t) = integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. Dirichlet_kernel n x * f(x+t)) / pi" (is "?lhs = ?rhs") proof - have ft: "(\x. f(x+t)) absolutely_integrable_on {-pi..pi}" using absolutely_integrable_periodic_offset [OF f, of t] periodic by simp have "?lhs = (\k=0..n. Fourier_coefficient (\x. f(x+t)) (2*k) * trigonometric_set (2*k) 0)" using Fourier_sum_offset_unpaired assms atMost_atLeast0 by auto also have "\ = Fourier_coefficient (\x. f(x+t)) 0 / sqrt (2 * pi) + (\k = Suc 0..n. Fourier_coefficient (\x. f(x+t)) (2*k) / sqrt pi)" by (simp add: sum.atLeast_Suc_atMost trigonometric_set_def) also have "\ = (LINT x|lebesgue_on {-pi..pi}. f(x+t)) / (2 * pi) + (\k = Suc 0..n. (LINT x|lebesgue_on {-pi..pi}. cos (real k * x) * f(x+t)) / pi)" by (simp add: Fourier_coefficient_def orthonormal_coeff_def trigonometric_set_def l2product_def) also have "\ = LINT x|lebesgue_on {-pi..pi}. f(x+t) / (2 * pi) + (\k = Suc 0..n. (cos (real k * x) * f(x+t)) / pi)" using Fourier_products_integrable_cos [OF ft] absolutely_integrable_imp_integrable [OF ft] by simp also have "\ = (LINT x|lebesgue_on {-pi..pi}. f(x+t) / 2 + (\k = Suc 0..n. cos (real k * x) * f(x+t))) / pi" by (simp add: divide_simps sum_distrib_right mult.assoc) also have "\ = ?rhs" proof - have "LINT x|lebesgue_on {-pi..pi}. f(x + t) / 2 + (\k = Suc 0..n. cos (real k * x) * f(x + t)) = LINT x|lebesgue_on {-pi..pi}. Dirichlet_kernel n x * f(x + t)" proof - have eq: "f(x+t) / 2 + (\k = Suc 0..n. cos (real k * x) * f(x + t)) = Dirichlet_kernel n x * f(x + t)" if "- pi \ x" "x \ pi" for x proof (cases "x = 0") case False then have "sin (x/2) \ 0" using that by (auto simp: sin_zero_iff) then have "f(x + t) * (1/2 + (\k = Suc 0..n. cos(real k * x))) = f(x + t) * sin((real n + 1/2) * x) / 2 / sin(x/2)" using cosine_sum_lemma [of x n] by (simp add: divide_simps) then show ?thesis by (simp add: Dirichlet_kernel_def False field_simps sum_distrib_left) qed (simp add: Dirichlet_kernel_def algebra_simps) show ?thesis by (rule Bochner_Integration.integral_cong [OF refl]) (simp add: eq) qed then show ?thesis by simp qed finally show ?thesis . qed lemma Fourier_sum_limit_Dirichlet_kernel: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" shows "((\n. (\k\n. Fourier_coefficient f k * trigonometric_set k t)) \ l) \ (\n. LINT x|lebesgue_on {-pi..pi}. Dirichlet_kernel n x * f(x + t)) \ pi * l" (is "?lhs = ?rhs") proof - have "?lhs \ (\n. (LINT x|lebesgue_on {-pi..pi}. Dirichlet_kernel n x * f(x + t)) / pi) \ l" by (simp add: Fourier_sum_limit_pair [OF f, symmetric] Fourier_sum_offset_Dirichlet_kernel assms) also have "\ \ (\k. (LINT x|lebesgue_on {-pi..pi}. Dirichlet_kernel k x * f(x + t)) * inverse pi) \ pi * l * inverse pi" by (auto simp: divide_simps) also have "\ \ ?rhs" using tendsto_mult_right_iff [of "inverse pi", symmetric] by auto finally show ?thesis . qed subsection\A directly deduced sufficient condition for convergence at a point\ lemma simple_Fourier_convergence_periodic: assumes f: "f absolutely_integrable_on {-pi..pi}" and ft: "(\x. (f(x+t) - f t) / sin(x/2)) absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" shows "(\n. (\k\n. Fourier_coefficient f k * trigonometric_set k t)) \ f t" proof - let ?\ = "\n. \k\n. Fourier_coefficient (\x. f x - f t) k * trigonometric_set k t" let ?\ = "\n. LINT x|lebesgue_on {-pi..pi}. Dirichlet_kernel n x * (f(x + t) - f t)" have fft: "(\x. f x - f t) absolutely_integrable_on {-pi..pi}" by (simp add: f absolutely_integrable_measurable_real) have fxt: "(\x. f(x + t)) absolutely_integrable_on {-pi..pi}" using absolutely_integrable_periodic_offset assms by auto have *: "?\ \ 0 \ ?\ \ 0" by (simp add: Fourier_sum_limit_Dirichlet_kernel fft periodic) moreover have "(\n. (\k\n. Fourier_coefficient f k * trigonometric_set k t) - f t) \ 0" if "?\ \ 0" proof (rule Lim_transform_eventually [OF that eventually_sequentiallyI]) show "(\k\n. Fourier_coefficient (\x. f x - f t) k * trigonometric_set k t) = (\k\n. Fourier_coefficient f k * trigonometric_set k t) - f t" if "Suc 0 \ n" for n proof - have "(\k = Suc 0..n. Fourier_coefficient (\x. f x - f t) k * trigonometric_set k t) = (\k = Suc 0..n. Fourier_coefficient f k * trigonometric_set k t)" proof (rule sum.cong [OF refl]) fix k assume k: "k \ {Suc 0..n}" then have [simp]: "integral\<^sup>L (lebesgue_on {-pi..pi}) (trigonometric_set k) = 0" by (auto simp: trigonometric_set_def) show "Fourier_coefficient (\x. f x - f t) k * trigonometric_set k t = Fourier_coefficient f k * trigonometric_set k t" using k unfolding Fourier_coefficient_def orthonormal_coeff_def l2product_def by (simp add: right_diff_distrib f absolutely_integrable_measurable_real trigonometric_set_mul_integrable) qed moreover have "Fourier_coefficient (\x. f x - f t) 0 * trigonometric_set 0 t = Fourier_coefficient f 0 * trigonometric_set 0 t - f t" using f unfolding Fourier_coefficient_def orthonormal_coeff_def l2product_def by (auto simp: divide_simps trigonometric_set absolutely_integrable_imp_integrable measure_restrict_space) ultimately show ?thesis by (simp add: sum.atLeast_Suc_atMost atMost_atLeast0) qed qed moreover have "?\ \ 0" proof - have eq: "\n. ?\ n = integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. sin((n + 1/2) * x) * ((f(x+t) - f t) / sin(x/2) / 2))" unfolding Dirichlet_kernel_def by (rule Bochner_Integration.integral_cong [OF refl]) (auto simp: divide_simps sin_zero_iff) show ?thesis unfolding eq by (intro ft set_integrable_divide Riemann_lebesgue_sin_half) qed ultimately show ?thesis by (simp add: LIM_zero_cancel) qed subsection\A more natural sufficient Hoelder condition at a point\ lemma bounded_inverse_sin_half: assumes "d > 0" obtains B where "B>0" "\x. x \ ({-pi..pi} - {-d<.. \inverse (sin (x/2))\ \ B" proof - have "bounded ((\x. inverse (sin (x/2))) ` ({-pi..pi} - {- d<..x \ {-pi..pi}; x \ 0\ \ sin (x/2) \ 0" for x using \0 < d\ by (auto simp: sin_zero_iff) then show "continuous_on ({-pi..pi} - {-d<..x. inverse (sin (x/2)))" using \0 < d\ by (intro continuous_intros) auto show "compact ({-pi..pi} - {-d<.. 0" "a > 0" and ft: "\x. \x-t\ < d \ \f x - f t\ \ M * \x-t\ powr a" and periodic: "\x. f(x + 2*pi) = f x" shows "(\n. (\k\n. Fourier_coefficient f k * trigonometric_set k t)) \ f t" proof (rule simple_Fourier_convergence_periodic [OF f]) have half: "((\x. sin(x/2)) has_real_derivative 1/2) (at 0)" by (rule derivative_eq_intros refl | force)+ then obtain e0::real where "e0 > 0" and e0: "\x. \x \ 0; \x\ < e0\ \ \sin (x/2) / x - 1/2\ < 1/4" apply (simp add: DERIV_def Lim_at dist_real_def) apply (drule_tac x="1/4" in spec, auto) done obtain e where "e > 0" and e: "\x. \x\ < e \ \(f(x+t) - f t) / sin (x/2)\ \ 4 * (\M\ * \x\ powr (a-1))" proof show "min d e0 > 0" using \d > 0\ \e0 > 0\ by auto next fix x assume x: "\x\ < min d e0" show "\(f(x + t) - f t) / sin (x/2)\ \ 4 * (\M\ * \x\ powr (a - 1))" proof (cases "sin(x/2) = 0 \ x=0") case False have eq: "\(f(x + t) - f t) / sin (x/2)\ = \inverse (sin (x/2) / x)\ * (\f(x + t) - f t\ / \x\)" by simp show ?thesis unfolding eq proof (rule mult_mono) have "\sin (x/2) / x - 1/2\ < 1/4" using e0 [of x] x False by force then have "1/4 \ \sin (x/2) / x\" by (simp add: abs_if field_simps split: if_split_asm) then show "\inverse (sin (x/2) / x)\ \ 4" using False by (simp add: field_simps) have "\f(x + t) - f t\ / \x\ \ M * \x + t - t\ powr (a - 1)" using ft [of "x+t"] x by (auto simp: divide_simps algebra_simps Transcendental.powr_mult_base) also have "\ \ \M\ * \x\ powr (a - 1)" by (simp add: mult_mono) finally show "\f(x + t) - f t\ / \x\ \ \M\ * \x\ powr (a - 1)" . qed auto qed auto qed obtain B where "B>0" and B: "\x. x \ ({-pi..pi} - {- e<.. \inverse (sin (x/2))\ \ B" using bounded_inverse_sin_half [OF \e > 0\] by metis let ?g = "\x. max (B * \f(x + t) - f t\) ((4 * \M\) * \x\ powr (a - 1))" show "(\x. (f(x + t) - f t) / sin (x/2)) absolutely_integrable_on {-pi..pi}" proof (rule measurable_bounded_by_integrable_imp_absolutely_integrable) have fxt: "(\x. f(x + t)) absolutely_integrable_on {-pi..pi}" using absolutely_integrable_periodic_offset assms by auto show "(\x. (f(x + t) - f t) / sin (x/2)) \ borel_measurable (lebesgue_on {-pi..pi})" proof (intro measurable) show "(\x. f(x + t)) \ borel_measurable (lebesgue_on {-pi..pi})" using absolutely_integrable_on_def fxt integrable_imp_measurable by blast show "(\x. sin (x/2)) \ borel_measurable (lebesgue_on {-pi..pi})" by (intro continuous_imp_measurable_on_sets_lebesgue continuous_intros) auto qed auto have "(\x. f(x + t) - f t) absolutely_integrable_on {-pi..pi}" by (intro set_integral_diff fxt) (simp add: set_integrable_def) moreover have "(\x. \x\ powr (a - 1)) absolutely_integrable_on {-pi..pi}" proof - have "((\x. x powr (a - 1)) has_integral inverse a * pi powr a - inverse a * 0 powr a) {0..pi}" proof (rule fundamental_theorem_of_calculus_interior) show "continuous_on {0..pi} (\x. inverse a * x powr a)" using \a > 0\ by (intro continuous_on_powr' continuous_intros) auto show "((\b. inverse a * b powr a) has_vector_derivative x powr (a - 1)) (at x)" if "x \ {0<..a > 0\ apply (simp flip: has_real_derivative_iff_has_vector_derivative) apply (rule derivative_eq_intros | simp)+ done qed auto then have int: "(\x. x powr (a - 1)) integrable_on {0..pi}" by blast show ?thesis proof (rule nonnegative_absolutely_integrable_1) show "(\x. \x\ powr (a - 1)) integrable_on {-pi..pi}" proof (rule Henstock_Kurzweil_Integration.integrable_combine) show "(\x. \x\ powr (a - 1)) integrable_on {0..pi}" using int integrable_eq by fastforce then show "(\x. \x\ powr (a - 1)) integrable_on {- pi..0}" using Henstock_Kurzweil_Integration.integrable_reflect_real by fastforce qed auto qed auto qed ultimately show "?g integrable_on {-pi..pi}" apply (intro set_lebesgue_integral_eq_integral absolutely_integrable_max_1 absolutely_integrable_bounded_measurable_product_real set_integrable_abs measurable) apply (auto simp: image_constant_conv) done show "norm ((f(x + t) - f t) / sin (x/2)) \ ?g x" if "x \ {-pi..pi}" for x proof (cases "\x\ < e") case True then show ?thesis using e [OF True] by (simp add: max_def) next case False then have "\inverse (sin (x/2))\ \ B" using B that by force then have "\inverse (sin (x/2))\ * \f(x + t) - f t\ \ B * \f(x + t) - f t\" by (simp add: mult_right_mono) then have "\f(x + t) - f t\ / \sin (x/2)\ \ B * \f(x + t) - f t\" by (simp add: divide_simps split: if_split_asm) then show ?thesis by auto qed qed auto qed (auto simp: periodic) text\In particular, a Lipschitz condition at the point\ corollary Lipschitz_Fourier_convergence_periodic: assumes f: "f absolutely_integrable_on {-pi..pi}" and "d > 0" and ft: "\x. \x-t\ < d \ \f x - f t\ \ M * \x-t\" and periodic: "\x. f(x + 2*pi) = f x" shows "(\n. (\k\n. Fourier_coefficient f k * trigonometric_set k t)) \ f t" using Hoelder_Fourier_convergence_periodic [OF f \d > 0\, of 1] assms by (fastforce simp add:) text\In particular, if left and right derivatives both exist\ proposition bi_differentiable_Fourier_convergence_periodic: assumes f: "f absolutely_integrable_on {-pi..pi}" and f_lt: "f differentiable at_left t" and f_gt: "f differentiable at_right t" and periodic: "\x. f(x + 2*pi) = f x" shows "(\n. (\k\n. Fourier_coefficient f k * trigonometric_set k t)) \ f t" proof - obtain D_lt where D_lt: "\e. e > 0 \ \d>0. \x dist x t < d \ dist ((f x - f t) / (x - t)) D_lt < e" using f_lt unfolding has_field_derivative_iff real_differentiable_def Lim_within by (meson lessThan_iff) then obtain d_lt where "d_lt > 0" and d_lt: "\x. \x < t; 0 < \x - t\; \x - t\ < d_lt\ \ \(f x - f t) / (x - t) - D_lt\ < 1" unfolding dist_real_def by (metis dist_commute dist_real_def zero_less_one) obtain D_gt where D_gt: "\e. e > 0 \ \d>0. \x>t. 0 < dist x t \ dist x t < d \ dist ((f x - f t) / (x - t)) D_gt < e" using f_gt unfolding has_field_derivative_iff real_differentiable_def Lim_within by (meson greaterThan_iff) then obtain d_gt where "d_gt > 0" and d_gt: "\x. \x > t; 0 < \x - t\; \x - t\ < d_gt\ \ \(f x - f t) / (x - t) - D_gt\ < 1" unfolding dist_real_def by (metis dist_commute dist_real_def zero_less_one) show ?thesis proof (rule Lipschitz_Fourier_convergence_periodic [OF f]) fix x assume "\x - t\ < min d_lt d_gt" then have *: "\x - t\ < d_lt" "\x - t\ < d_gt" by auto consider "xt" by linarith then show "\f x - f t\ \ (\D_lt\ + \D_gt\ + 1) * \x - t\" proof cases case 1 then have "\(f x - f t) / (x - t) - D_lt\ < 1" using d_lt [OF 1] * by auto then have "\(f x - f t) / (x - t)\ - \D_lt\ < 1" by linarith then have "\f x - f t\ \ (\D_lt\ + 1) * \x - t\" by (simp add: comm_semiring_class.distrib divide_simps split: if_split_asm) also have "\ \ (\D_lt\ + \D_gt\ + 1) * \x - t\" by (simp add: mult_right_mono) finally show ?thesis . next case 3 then have "\(f x - f t) / (x - t) - D_gt\ < 1" using d_gt [OF 3] * by auto then have "\(f x - f t) / (x - t)\ - \D_gt\ < 1" by linarith then have "\f x - f t\ \ (\D_gt\ + 1) * \x - t\" by (simp add: comm_semiring_class.distrib divide_simps split: if_split_asm) also have "\ \ (\D_lt\ + \D_gt\ + 1) * \x - t\" by (simp add: mult_right_mono) finally show ?thesis . qed auto qed (auto simp: \0 < d_gt\ \0 < d_lt\ periodic) qed text\And in particular at points where the function is differentiable\ lemma differentiable_Fourier_convergence_periodic: assumes f: "f absolutely_integrable_on {-pi..pi}" and fdif: "f differentiable (at t)" and periodic: "\x. f(x + 2*pi) = f x" shows "(\n. (\k\n. Fourier_coefficient f k * trigonometric_set k t)) \ f t" by (rule bi_differentiable_Fourier_convergence_periodic) (auto simp: differentiable_at_withinI assms) text\Use reflection to halve the region of integration\ lemma absolutely_integrable_mult_Dirichlet_kernel_reflected: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" shows "(\x. Dirichlet_kernel n x * f(t+x)) absolutely_integrable_on {-pi..pi}" "(\x. Dirichlet_kernel n x * f(t-x)) absolutely_integrable_on {-pi..pi}" "(\x. Dirichlet_kernel n x * c) absolutely_integrable_on {-pi..pi}" proof - show "(\x. Dirichlet_kernel n x * f(t+x)) absolutely_integrable_on {-pi..pi}" apply (rule absolutely_integrable_mult_Dirichlet_kernel) using absolutely_integrable_periodic_offset [OF f, of t] periodic apply simp done then show "(\x. Dirichlet_kernel n x * f(t-x)) absolutely_integrable_on {-pi..pi}" by (subst absolutely_integrable_reflect_real [symmetric]) simp show "(\x. Dirichlet_kernel n x * c) absolutely_integrable_on {-pi..pi}" by (simp add: absolutely_integrable_measurable_real borel_measurable_integrable integrable_Dirichlet_kernel) qed lemma absolutely_integrable_mult_Dirichlet_kernel_reflected_part: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" and "d \ pi" shows "(\x. Dirichlet_kernel n x * f(t+x)) absolutely_integrable_on {0..d}" "(\x. Dirichlet_kernel n x * f(t-x)) absolutely_integrable_on {0..d}" "(\x. Dirichlet_kernel n x * c) absolutely_integrable_on {0..d}" using absolutely_integrable_mult_Dirichlet_kernel_reflected [OF f periodic, of n] \d \ pi\ by (force intro: absolutely_integrable_on_subinterval)+ lemma absolutely_integrable_mult_Dirichlet_kernel_reflected_part2: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" and "d \ pi" shows "(\x. Dirichlet_kernel n x * (f(t+x) + f(t-x))) absolutely_integrable_on {0..d}" "(\x. Dirichlet_kernel n x * ((f(t+x) + f(t-x)) - c)) absolutely_integrable_on {0..d}" using absolutely_integrable_mult_Dirichlet_kernel_reflected_part assms by (simp_all add: distrib_left right_diff_distrib) lemma integral_reflect_and_add: fixes f :: "real \ 'b::euclidean_space" assumes "integrable (lebesgue_on {-a..a}) f" shows "integral\<^sup>L (lebesgue_on {-a..a}) f = integral\<^sup>L (lebesgue_on {0..a}) (\x. f x + f(-x))" proof (cases "a \ 0") case True have f: "integrable (lebesgue_on {0..a}) f" "integrable (lebesgue_on {-a..0}) f" using integrable_subinterval assms by fastforce+ then have fm: "integrable (lebesgue_on {0..a}) (\x. f(-x))" using integrable_reflect_real by fastforce have "integral\<^sup>L (lebesgue_on {-a..a}) f = integral\<^sup>L (lebesgue_on {-a..0}) f + integral\<^sup>L (lebesgue_on {0..a}) f" by (simp add: True assms integral_combine) also have "\ = integral\<^sup>L (lebesgue_on {0..a}) (\x. f(-x)) + integral\<^sup>L (lebesgue_on {0..a}) f" by (metis (no_types) add.inverse_inverse add.inverse_neutral integral_reflect_real) also have "\ = integral\<^sup>L (lebesgue_on {0..a}) (\x. f x + f(-x))" by (simp add: fm f) finally show ?thesis . qed (use assms in auto) lemma Fourier_sum_offset_Dirichlet_kernel_half: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" shows "(\k\2*n. Fourier_coefficient f k * trigonometric_set k t) - l = (LINT x|lebesgue_on {0..pi}. Dirichlet_kernel n x * (f(t+x) + f(t-x) - 2*l)) / pi" proof - have fxt: "(\x. f(x + t)) absolutely_integrable_on {-pi..pi}" using absolutely_integrable_periodic_offset assms by auto have int: "integrable (lebesgue_on {0..pi}) (Dirichlet_kernel n)" using not_integrable_integral_eq by fastforce have "LINT x|lebesgue_on {-pi..pi}. Dirichlet_kernel n x * f(x + t) = LINT x|lebesgue_on {0..pi}. Dirichlet_kernel n x * f(x + t) + Dirichlet_kernel n (- x) * f(- x + t)" by (simp add: integral_reflect_and_add absolutely_integrable_imp_integrable absolutely_integrable_mult_Dirichlet_kernel fxt) also have "\ = (LINT x|lebesgue_on {0..pi}. Dirichlet_kernel n x * (f(t+x) + f(t-x) - 2*l)) + pi * l" using absolutely_integrable_mult_Dirichlet_kernel_reflected_part [OF f periodic order_refl, of n t] apply (simp add: algebra_simps absolutely_integrable_imp_integrable int) done finally show ?thesis by (simp add: Fourier_sum_offset_Dirichlet_kernel assms divide_simps) qed lemma Fourier_sum_limit_Dirichlet_kernel_half: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" shows "(\n. (\k\n. Fourier_coefficient f k * trigonometric_set k t)) \ l \ (\n. (LINT x|lebesgue_on {0..pi}. Dirichlet_kernel n x * (f(t+x) + f(t-x) - 2*l))) \ 0" apply (simp flip: Fourier_sum_limit_pair [OF f]) apply (simp add: Lim_null [where l=l] Fourier_sum_offset_Dirichlet_kernel_half assms) done subsection\Localization principle: convergence only depends on values "nearby"\ proposition Riemann_localization_integral: assumes f: "f absolutely_integrable_on {-pi..pi}" and g: "g absolutely_integrable_on {-pi..pi}" and "d > 0" and d: "\x. \x\ < d \ f x = g x" shows "(\n. integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. Dirichlet_kernel n x * f x) - integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. Dirichlet_kernel n x * g x)) \ 0" (is "?a \ 0") proof - let ?f = "\x. (if \x\ < d then 0 else f x - g x) / sin(x/2) / 2" have eq: "?a n = integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. sin((n+1/2) * x) * ?f x)" for n proof - have "?a n = integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. Dirichlet_kernel n x * (if \x\ < d then 0 else f x - g x))" apply (simp add: absolutely_integrable_imp_integrable absolutely_integrable_mult_Dirichlet_kernel f g flip: Bochner_Integration.integral_diff) apply (auto simp: d algebra_simps intro: Bochner_Integration.integral_cong) done also have "\ = integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. sin((n+1/2) * x) * ?f x)" using \d > 0\ by (auto simp: Dirichlet_kernel_def intro: Bochner_Integration.integral_cong) finally show ?thesis . qed show ?thesis unfolding eq proof (rule Riemann_lebesgue_sin_half) obtain B where "B>0" and B: "\x. x \ ({-pi..pi} - {-d<.. \inverse (sin (x/2))\ \ B" using bounded_inverse_sin_half [OF \d > 0\] by metis have "(\x. (if \x\ < d then 0 else f x - g x) / sin (x/2)) absolutely_integrable_on {-pi..pi}" proof (rule measurable_bounded_by_integrable_imp_absolutely_integrable) show "(\x. (if \x\ < d then 0 else f x - g x) / sin (x/2)) \ borel_measurable (lebesgue_on {-pi..pi})" proof (intro measurable) show "f \ borel_measurable (lebesgue_on {-pi..pi})" "g \ borel_measurable (lebesgue_on {-pi..pi})" using absolutely_integrable_on_def f g integrable_imp_measurable by blast+ show "(\x. x) \ borel_measurable (lebesgue_on {-pi..pi})" "(\x. sin (x/2)) \ borel_measurable (lebesgue_on {-pi..pi})" by (intro continuous_intros continuous_imp_measurable_on_sets_lebesgue | force)+ qed auto have "(\x. B * \f x - g x\) absolutely_integrable_on {-pi..pi}" using \B > 0\ by (simp add: f g set_integrable_abs) then show "(\x. B * \f x - g x\) integrable_on {-pi..pi}" using absolutely_integrable_on_def by blast next fix x assume x: "x \ {-pi..pi}" then have [simp]: "sin (x/2) = 0 \ x=0" using \0 < d\ by (force simp: sin_zero_iff) show "norm ((if \x\ < d then 0 else f x - g x) / sin (x/2)) \ B * \f x - g x\" proof (cases "\x\ < d") case False then have "1 \ B * \sin (x/2)\" using \B > 0\ B [of x] x \d > 0\ by (auto simp: abs_less_iff divide_simps split: if_split_asm) then have "1 * \f x - g x\ \ B * \sin (x/2)\ * \f x - g x\" by (metis (full_types) abs_ge_zero mult.commute mult_left_mono) then show ?thesis using False by (auto simp: divide_simps algebra_simps) qed (simp add: d) qed auto then show "(\x. (if \x\ < d then 0 else f x - g x) / sin (x/2) / 2) absolutely_integrable_on {-pi..pi}" using set_integrable_divide by blast qed qed lemma Riemann_localization_integral_range: assumes f: "f absolutely_integrable_on {-pi..pi}" and "0 < d" "d \ pi" shows "(\n. integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. Dirichlet_kernel n x * f x) - integral\<^sup>L (lebesgue_on {-d..d}) (\x. Dirichlet_kernel n x * f x)) \ 0" proof - have *: "(\n. (LINT x|lebesgue_on {-pi..pi}. Dirichlet_kernel n x * f x) - (LINT x|lebesgue_on {-pi..pi}. Dirichlet_kernel n x * (if x \ {-d..d} then f x else 0))) \ 0" proof (rule Riemann_localization_integral [OF f _ \0 < d\]) have "f absolutely_integrable_on {-d..d}" using f absolutely_integrable_on_subinterval \d \ pi\ by fastforce moreover have "(\x. if - pi \ x \ x \ pi then if x \ {-d..d} then f x else 0 else 0) = (\x. if x \ {-d..d} then f x else 0)" using \d \ pi\ by force ultimately show "(\x. if x \ {-d..d} then f x else 0) absolutely_integrable_on {-pi..pi}" apply (subst absolutely_integrable_restrict_UNIV [symmetric]) apply (simp flip: absolutely_integrable_restrict_UNIV [of "{-d..d}" f]) done qed auto then show ?thesis using integral_restrict [of "{-d..d}" "{-pi..pi}" "\x. Dirichlet_kernel _ x * f x"] assms by (simp add: if_distrib cong: if_cong) qed lemma Riemann_localization: assumes f: "f absolutely_integrable_on {-pi..pi}" and g: "g absolutely_integrable_on {-pi..pi}" and perf: "\x. f(x + 2*pi) = f x" and perg: "\x. g(x + 2*pi) = g x" and "d > 0" and d: "\x. \x-t\ < d \ f x = g x" shows "(\n. \k\n. Fourier_coefficient f k * trigonometric_set k t) \ c \ (\n. \k\n. Fourier_coefficient g k * trigonometric_set k t) \ c" proof - have "(\n. LINT x|lebesgue_on {-pi..pi}. Dirichlet_kernel n x * f(x + t)) \ pi * c \ (\n. LINT x|lebesgue_on {-pi..pi}. Dirichlet_kernel n x * g (x + t)) \ pi * c" proof (intro Lim_transform_eq Riemann_localization_integral) show "(\x. f(x + t)) absolutely_integrable_on {-pi..pi}" "(\x. g (x + t)) absolutely_integrable_on {-pi..pi}" using assms by (auto intro: absolutely_integrable_periodic_offset) qed (use assms in auto) then show ?thesis by (simp add: Fourier_sum_limit_Dirichlet_kernel assms) qed subsection\Localize the earlier integral\ lemma Riemann_localization_integral_range_half: assumes f: "f absolutely_integrable_on {-pi..pi}" and "0 < d" "d \ pi" shows "(\n. (LINT x|lebesgue_on {0..pi}. Dirichlet_kernel n x * (f x + f(-x))) - (LINT x|lebesgue_on {0..d}. Dirichlet_kernel n x * (f x + f(-x)))) \ 0" proof - have *: "(\x. Dirichlet_kernel n x * f x) absolutely_integrable_on {-d..d}" for n by (metis (full_types) absolutely_integrable_mult_Dirichlet_kernel absolutely_integrable_on_subinterval \d \ pi\ atLeastatMost_subset_iff f neg_le_iff_le) show ?thesis using absolutely_integrable_mult_Dirichlet_kernel [OF f] using Riemann_localization_integral_range [OF assms] apply (simp add: "*" absolutely_integrable_imp_integrable integral_reflect_and_add) apply (simp add: algebra_simps) done qed lemma Fourier_sum_limit_Dirichlet_kernel_part: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" and d: "0 < d" "d \ pi" shows "(\n. \k\n. Fourier_coefficient f k * trigonometric_set k t) \ l \ (\n. (LINT x|lebesgue_on {0..d}. Dirichlet_kernel n x * ((f(t+x) + f(t-x)) - 2*l))) \ 0" proof - have "(\x. f(t+x)) absolutely_integrable_on {-pi..pi}" using absolutely_integrable_periodic_offset [OF f, of t] periodic by simp then have int: "(\x. f(t+x) - l) absolutely_integrable_on {-pi..pi}" by auto have "(\n. LINT x|lebesgue_on {0..pi}. Dirichlet_kernel n x * (f(t+x) + f(t-x) - 2*l)) \ 0 \ (\n. LINT x|lebesgue_on {0..d}. Dirichlet_kernel n x * (f(t+x) + f(t-x) - 2*l)) \ 0" by (rule Lim_transform_eq) (use Riemann_localization_integral_range_half [OF int d] in auto) then show ?thesis by (simp add: Fourier_sum_limit_Dirichlet_kernel_half assms) qed subsection\Make a harmless simplifying tweak to the Dirichlet kernel\ lemma inte_Dirichlet_kernel_mul_expand: assumes f: "f \ borel_measurable (lebesgue_on S)" and S: "S \ sets lebesgue" shows "(LINT x|lebesgue_on S. Dirichlet_kernel n x * f x = LINT x|lebesgue_on S. sin((n+1/2) * x) * f x / (2 * sin(x/2))) \ (integrable (lebesgue_on S) (\x. Dirichlet_kernel n x * f x) \ integrable (lebesgue_on S) (\x. sin((n+1/2) * x) * f x / (2 * sin(x/2))))" proof (cases "0 \ S") case True have *: "{x. x = 0 \ x \ S} \ sets (lebesgue_on S)" using True by (simp add: S sets_restrict_space_iff cong: conj_cong) have bm1: "(\x. Dirichlet_kernel n x * f x) \ borel_measurable (lebesgue_on S)" unfolding Dirichlet_kernel_def by (force intro: * assms measurable continuous_imp_measurable_on_sets_lebesgue continuous_intros) have bm2: "(\x. sin ((n + 1/2) * x) * f x / (2 * sin (x/2))) \ borel_measurable (lebesgue_on S)" by (intro assms measurable continuous_imp_measurable_on_sets_lebesgue continuous_intros) auto have 0: "{0} \ null_sets (lebesgue_on S)" using True by (simp add: S null_sets_restrict_space) show ?thesis apply (intro conjI integral_cong_AE integrable_cong_AE bm1 bm2 AE_I' [OF 0]) unfolding Dirichlet_kernel_def by auto next case False then show ?thesis unfolding Dirichlet_kernel_def by (auto intro!: Bochner_Integration.integral_cong Bochner_Integration.integrable_cong) qed lemma assumes f: "f \ borel_measurable (lebesgue_on S)" and S: "S \ sets lebesgue" shows integral_Dirichlet_kernel_mul_expand: "(LINT x|lebesgue_on S. Dirichlet_kernel n x * f x) = (LINT x|lebesgue_on S. sin((n+1/2) * x) * f x / (2 * sin(x/2)))" (is "?th1") and integrable_Dirichlet_kernel_mul_expand: "integrable (lebesgue_on S) (\x. Dirichlet_kernel n x * f x) \ integrable (lebesgue_on S) (\x. sin((n+1/2) * x) * f x / (2 * sin(x/2)))" (is "?th2") using inte_Dirichlet_kernel_mul_expand [OF assms] by auto proposition Fourier_sum_limit_sine_part: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" and d: "0 < d" "d \ pi" shows "(\n. (\k\n. Fourier_coefficient f k * trigonometric_set k t)) \ l \ (\n. LINT x|lebesgue_on {0..d}. sin((n + 1/2) * x) * ((f(t+x) + f(t-x) - 2*l) / x)) \ 0" (is "?lhs \ ?\ \ 0") proof - have ftx: "(\x. f(t+x)) absolutely_integrable_on {-pi..pi}" using absolutely_integrable_periodic_offset assms by auto then have ftmx: "(\x. f(t-x)) absolutely_integrable_on {-pi..pi}" by (simp flip: absolutely_integrable_reflect_real [where f= "(\x. f(t+x))"]) have fbm: "(\x. f(t+x) + f(t-x) - 2*l) absolutely_integrable_on {-pi..pi}" by (force intro: ftmx ftx) let ?\ = "\n. LINT x|lebesgue_on {0..d}. Dirichlet_kernel n x * ((f(t+x) + f(t-x)) - 2*l)" have "?lhs \ ?\ \ 0" by (intro Fourier_sum_limit_Dirichlet_kernel_part assms) moreover have "?\ \ 0 \ ?\ \ 0" proof (rule Lim_transform_eq [OF Lim_transform_eventually]) let ?f = "\n. LINT x|lebesgue_on {0..d}. sin((real n + 1/2) * x) * (1 / (2 * sin(x/2)) - 1/x) * (f(t+x) + f(t-x) - 2*l)" have "?f n = (LINT x|lebesgue_on {-pi..pi}. sin ((n + 1/2) * x) * ((if x \ {0..d} then 1 / (2 * sin (x/2)) - 1/x else 0) * (f(t+x) + f(t-x) - 2*l)))" for n using d by (simp add: integral_restrict if_distrib [of "\u. _ * (u * _)"] mult.assoc flip: atLeastAtMost_iff cong: if_cong) moreover have "\ \ 0" proof (intro Riemann_lebesgue_sin_half absolutely_integrable_bounded_measurable_product_real) have "(\x. 1 / (2 * sin(x/2)) - 1/x) \ borel_measurable (lebesgue_on {0..d})" by (intro measurable continuous_imp_measurable_on_sets_lebesgue continuous_intros) auto then show "(\x. if x \ {0..d} then 1 / (2 * sin(x/2)) - 1/x else 0) \ borel_measurable (lebesgue_on {-pi..pi})" using d by (simp add: borel_measurable_if_lebesgue_on flip: atLeastAtMost_iff) let ?g = "\x::real. 1 / (2 * sin(x/2)) - 1/x" have limg: "(?g \ ?g a) (at a within {0..d})" \\thanks to Manuel Eberl\ if a: "0 \ a" "a \ d" for a proof - have "(?g \ 0) (at_right 0)" and "(?g \ ?g d) (at_left d)" using d sin_gt_zero[of "d/2"] by (real_asymp simp: field_simps)+ moreover have "(?g \ ?g a) (at a)" if "a > 0" using a that d sin_gt_zero[of "a/2"] that by (real_asymp simp: field_simps) ultimately show ?thesis using that d by (cases "a = 0 \ a = d") (auto simp: at_within_Icc_at at_within_Icc_at_right at_within_Icc_at_left) qed have "((\x. if x \ {0..d} then 1 / (2 * sin(x/2)) - 1/x else 0) ` {-pi..pi}) = ((\x. 1 / (2 * sin(x/2)) - 1/x) ` {0..d})" using d by (auto intro: image_eqI [where x=0]) moreover have "bounded \" by (intro compact_imp_bounded compact_continuous_image) (auto simp: continuous_on limg) ultimately show "bounded ((\x. if x \ {0..d} then 1 / (2 * sin(x/2)) - 1/x else 0) ` {-pi..pi})" by simp qed (auto simp: fbm) ultimately show "?f \ (0::real)" by simp show "\\<^sub>F n in sequentially. ?f n = ?\ n - ?\ n" proof (rule eventually_sequentiallyI [where c = "Suc 0"]) fix n assume "n \ Suc 0" have "integrable (lebesgue_on {0..d}) (\x. sin ((real n + 1/2) * x) * (f(t+x) + f(t-x) - 2*l) / (2 * sin(x/2)))" using d apply (subst integrable_Dirichlet_kernel_mul_expand [symmetric]) apply (intro assms absolutely_integrable_imp_borel_measurable absolutely_integrable_on_subinterval [OF fbm] absolutely_integrable_imp_integrable absolutely_integrable_mult_Dirichlet_kernel_reflected_part2 | force)+ done moreover have "integrable (lebesgue_on {0..d}) (\x. sin ((real n + 1/2) * x) * (f(t+x) + f(t-x) - 2*l) / x)" proof (rule integrable_cong_AE_imp) let ?g = "\x. Dirichlet_kernel n x * (2 * sin(x/2) / x * (f(t+x) + f(t-x) - 2*l))" have *: "\2 * sin (x/2) / x\ \ 1" for x::real using abs_sin_x_le_abs_x [of "x/2"] by (simp add: divide_simps mult.commute abs_mult) have "bounded ((\x. 1 + (x/2)\<^sup>2) ` {-pi..pi})" by (intro compact_imp_bounded compact_continuous_image continuous_intros) auto then have bo: "bounded ((\x. 2 * sin (x/2) / x) ` {-pi..pi})" using * unfolding bounded_real by blast show "integrable (lebesgue_on {0..d}) ?g" using d apply (intro absolutely_integrable_imp_integrable absolutely_integrable_on_subinterval [OF absolutely_integrable_mult_Dirichlet_kernel] absolutely_integrable_bounded_measurable_product_real bo fbm measurable continuous_imp_measurable_on_sets_lebesgue continuous_intros, auto) done show "(\x. sin ((n + 1/2) * x) * (f(t+x) + f(t-x) - 2*l) / x) \ borel_measurable (lebesgue_on {0..d})" using d apply (intro measurable absolutely_integrable_imp_borel_measurable absolutely_integrable_on_subinterval [OF ftx] absolutely_integrable_on_subinterval [OF ftmx] absolutely_integrable_continuous_real continuous_intros, auto) done have "Dirichlet_kernel n x * (2 * sin(x/2)) / x = sin ((real n + 1/2) * x) / x" if "0 < x" "x \ d" for x using that d by (simp add: Dirichlet_kernel_def divide_simps sin_zero_pi_iff) then show "AE x in lebesgue_on {0..d}. ?g x = sin ((real n + 1/2) * x) * (f(t+x) + f(t-x) - 2*l) / x" using d by (force intro: AE_I' [where N="{0}"]) qed ultimately have "?f n = (LINT x|lebesgue_on {0..d}. sin ((n + 1/2) * x) * (f(t+x) + f(t-x) - 2*l) / (2 * sin(x/2))) - (LINT x|lebesgue_on {0..d}. sin ((n + 1/2) * x) * (f(t+x) + f(t-x) - 2*l) / x)" by (simp add: right_diff_distrib [of _ _ "1/_"] left_diff_distrib) also have "\ = ?\ n - ?\ n" using d by (simp add: measurable_restrict_mono [OF absolutely_integrable_imp_borel_measurable [OF fbm]] integral_Dirichlet_kernel_mul_expand) finally show "?f n = ?\ n - ?\ n" . qed qed ultimately show ?thesis by simp qed subsection\Dini's test for the convergence of a Fourier series\ proposition Fourier_Dini_test: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" and int0d: "integrable (lebesgue_on {0..d}) (\x. \f(t+x) + f(t-x) - 2*l\ / x)" and "0 < d" shows "(\n. (\k\n. Fourier_coefficient f k * trigonometric_set k t)) \ l" proof - define \ where "\ \ \n x. sin((n + 1/2) * x) * ((f(t+x) + f(t-x) - 2*l) / x)" have "(\n. LINT x|lebesgue_on {0..pi}. \ n x) \ 0" unfolding lim_sequentially dist_real_def proof (intro allI impI) fix e :: real assume "e > 0" define \ where "\ \ \x. LINT x|lebesgue_on {0..x}. \f(t+x) + f(t-x) - 2*l\ / x" have [simp]: "\ 0 = 0" by (simp add: \_def integral_eq_zero_null_sets) have cont: "continuous_on {0..d} \" unfolding \_def using indefinite_integral_continuous_real int0d by blast with \d > 0\ have "\e>0. \da>0. \x'\{0..d}. \x' - 0\ < da \ \\ x' - \ 0\ < e" by (force simp: continuous_on_real_range dest: bspec [where x=0]) with \e > 0\ obtain k where "k > 0" and k: "\x'. \0 \ x'; x' < k; x' \ d\ \ \\ x'\ < e/2" by (drule_tac x="e/2" in spec) auto define \ where "\ \ min d (min (k/2) pi)" have e2: "\\ \\ < e/2" and "\ > 0" "\ \ pi" unfolding \_def using k \k > 0\ \d > 0\ by auto have ftx: "(\x. f(t+x)) absolutely_integrable_on {-pi..pi}" using absolutely_integrable_periodic_offset assms by auto then have ftmx: "(\x. f(t-x)) absolutely_integrable_on {-pi..pi}" by (simp flip: absolutely_integrable_reflect_real [where f= "(\x. f(t+x))"]) have 1: "\ n absolutely_integrable_on {0..\}" for n::nat unfolding \_def proof (rule absolutely_integrable_bounded_measurable_product_real) show "(\x. sin ((real n + 1/2) * x)) \ borel_measurable (lebesgue_on {0..\})" by (intro continuous_imp_measurable_on_sets_lebesgue continuous_intros) auto show "bounded ((\x. sin ((real n + 1/2) * x)) ` {0..\})" by (simp add: bounded_iff) (use abs_sin_le_one in blast) have "(\x. (f(t+x) + f(t-x) - 2*l) / x) \ borel_measurable (lebesgue_on {0..\})" using \d > 0\ unfolding \_def by (intro measurable absolutely_integrable_imp_borel_measurable absolutely_integrable_on_subinterval [OF ftx] absolutely_integrable_on_subinterval [OF ftmx] absolutely_integrable_continuous_real continuous_intros, auto) moreover have "integrable (lebesgue_on {0..\}) (norm \ (\x. (f(t+x) + f(t-x) - 2*l) / x))" proof (rule integrable_subinterval [of 0 d]) show "integrable (lebesgue_on {0..d}) (norm \ (\x. (f(t+x) + f(t-x) - 2*l) / x))" using int0d by (subst Bochner_Integration.integrable_cong) (auto simp: o_def) show "{0..\} \ {0..d}" using \d > 0\ by (auto simp: \_def) qed ultimately show "(\x. (f(t+x) + f(t-x) - 2*l) / x) absolutely_integrable_on {0..\}" by (auto simp: absolutely_integrable_measurable) qed auto have 2: "\ n absolutely_integrable_on {\..pi}" for n::nat unfolding \_def proof (rule absolutely_integrable_bounded_measurable_product_real) show "(\x. sin ((n + 1/2) * x)) \ borel_measurable (lebesgue_on {\..pi})" by (intro continuous_imp_measurable_on_sets_lebesgue continuous_intros) auto show "bounded ((\x. sin ((n + 1/2) * x)) ` {\..pi})" by (simp add: bounded_iff) (use abs_sin_le_one in blast) have "(\x. inverse x * (f(t+x) + f(t-x) - 2*l)) absolutely_integrable_on {\..pi}" proof (rule absolutely_integrable_bounded_measurable_product_real) have "(\x. 1/x) \ borel_measurable (lebesgue_on {\..pi})" by (auto simp: measurable_completion measurable_restrict_space1) then show "inverse \ borel_measurable (lebesgue_on {\..pi})" by (metis (no_types, lifting) inverse_eq_divide measurable_lebesgue_cong) have "\x\{\..pi}. inverse \x\ \ inverse \" using \0 < \\ by auto then show "bounded (inverse ` {\..pi})" by (auto simp: bounded_iff) show "(\x. f(t+x) + f(t-x) - 2*l) absolutely_integrable_on {\..pi}" proof (rule absolutely_integrable_on_subinterval) show "(\x. (f(t+x) + f(t-x) - 2*l)) absolutely_integrable_on {-pi..pi}" by (fast intro: ftx ftmx absolutely_integrable_on_const) show "{\..pi} \ {-pi..pi}" using \0 < \\ by auto qed qed auto then show "(\x. (f(t+x) + f(t-x) - 2*l) / x) absolutely_integrable_on {\..pi}" by (metis (no_types, lifting) divide_inverse mult.commute set_integrable_cong) qed auto have "(\x. f(t+x) - l) absolutely_integrable_on {-pi..pi}" using ftx by auto moreover have "bounded ((\x. 0) ` {x. \x\ < \})" using bounded_def by blast moreover have "bounded (inverse ` {x. \ \x\ < \})" using \\ > 0\ by (auto simp: divide_simps intro: boundedI [where B = "1/\"]) ultimately have "(\x. (if \x\ < \ then 0 else inverse x) * (if x \ {-pi..pi} then f(t+x) - l else 0)) absolutely_integrable_on UNIV" apply (intro absolutely_integrable_bounded_measurable_product_real measurable set_integral_diff) apply (auto simp: lebesgue_on_UNIV_eq measurable_completion simp flip: absolutely_integrable_restrict_UNIV [where S = "{-pi..pi}"]) done moreover have "(if x \ {-pi..pi} then if \x\ < \ then 0 else (f(t+x) - l) / x else 0) = (if \x\ < \ then 0 else inverse x) * (if x \ {-pi..pi} then f(t+x) - l else 0)" for x by (auto simp: divide_simps) ultimately have *: "(\x. if \x\ < \ then 0 else (f(t+x) - l) / x) absolutely_integrable_on {-pi..pi}" by (simp add: flip: absolutely_integrable_restrict_UNIV [where S = "{-pi..pi}"]) have "(\n. LINT x|lebesgue_on {0..pi}. sin ((n + 1/2) * x) * (if \x\ < \ then 0 else (f(t+x) - l) / x) + sin ((n + 1/2) * -x) * (if \x\ < \ then 0 else (f(t-x) - l) / -x)) \ 0" using Riemann_lebesgue_sin_half [OF *] * by (simp add: integral_reflect_and_add absolutely_integrable_imp_integrable absolutely_integrable_sin_product cong: if_cong) moreover have "sin ((n + 1/2) * x) * (if \x\ < \ then 0 else (f(t+x) - l) / x) + sin ((n + 1/2) * -x) * (if \x\ < \ then 0 else (f(t-x) - l) / -x) = (if \ \x\ < \ then \ n x else 0)" for x n by simp (auto simp: divide_simps algebra_simps \_def) ultimately have "(\n. LINT x|lebesgue_on {0..pi}. (if \ \x\ < \ then \ n x else 0)) \ 0" by simp moreover have "(if 0 \ x \ x \ pi then if \ \x\ < \ then \ n x else 0 else 0) = (if \ \ x \ x \ pi then \ n x else 0)" for x n using \\ > 0\ by auto ultimately have \: "(\n. LINT x|lebesgue_on {\..pi}. \ n x) \ 0" by (simp flip: Lebesgue_Measure.integral_restrict_UNIV) then obtain N::nat where N: "\n. n\N \ \LINT x|lebesgue_on {\..pi}. \ n x\ < e/2" unfolding lim_sequentially dist_real_def by (metis (full_types) \0 < e\ diff_zero half_gt_zero_iff) have "\integral\<^sup>L (lebesgue_on {0..pi}) (\ n)\ < e" if "n \ N" for n::nat proof - have int: "integrable (lebesgue_on {0..pi}) (\ (real n))" by (intro integrable_combine [of concl: 0 pi] absolutely_integrable_imp_integrable) (use \\ > 0\ \\ \ pi\ 1 2 in auto) then have "integral\<^sup>L (lebesgue_on {0..pi}) (\ n) = integral\<^sup>L (lebesgue_on {0..\}) (\ n) + integral\<^sup>L (lebesgue_on {\..pi}) (\ n)" by (rule integral_combine) (use \0 < \\ \\ \ pi\ in auto) moreover have "\integral\<^sup>L (lebesgue_on {0..\}) (\ n)\ \ LINT x|lebesgue_on {0..\}. \f(t + x) + f(t - x) - 2 * l\ / x" proof (rule integral_abs_bound_integral) show "integrable (lebesgue_on {0..\}) (\ (real n))" by (meson integrable_subinterval \\ \ pi\ int atLeastatMost_subset_iff order_refl) have "{0..\} \ {0..d}" by (auto simp: \_def) then show "integrable (lebesgue_on {0..\}) (\x. \f(t + x) + f(t - x) - 2 * l\ / x)" by (rule integrable_subinterval [OF int0d]) show "\\ (real n) x\ \ \f(t + x) + f(t - x) - 2 * l\ / x" if "x \ space (lebesgue_on {0..\})" for x using that apply (auto simp: \_def divide_simps abs_mult) by (simp add: mult.commute mult_left_le) qed ultimately have "\integral\<^sup>L (lebesgue_on {0..pi}) (\ n)\ < e/2 + e/2" using N [OF that] e2 unfolding \_def by linarith then show ?thesis by simp qed then show "\N. \n\N. \integral\<^sup>L (lebesgue_on {0..pi}) (\ (real n)) - 0\ < e" by force qed then show ?thesis unfolding \_def using Fourier_sum_limit_sine_part assms pi_gt_zero by blast qed subsection\Cesaro summability of Fourier series using Fejér kernel\ definition Fejer_kernel :: "nat \ real \ real" where "Fejer_kernel \ \n x. if n = 0 then 0 else (\r sin(x/2) = 0") case True have "(\rrx. Fejer_kernel n x * f x) absolutely_integrable_on {-pi..pi}" proof (rule absolutely_integrable_bounded_measurable_product_real) show "Fejer_kernel n \ borel_measurable (lebesgue_on {-pi..pi})" by (simp add: Fejer_kernel_continuous continuous_imp_measurable_on_sets_lebesgue) show "bounded (Fejer_kernel n ` {-pi..pi})" using Fejer_kernel_continuous compact_continuous_image compact_imp_bounded by blast qed (use assms in auto) lemma absolutely_integrable_mult_Fejer_kernel_reflected1: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" shows "(\x. Fejer_kernel n x * f(t + x)) absolutely_integrable_on {-pi..pi}" using assms by (force intro: absolutely_integrable_mult_Fejer_kernel absolutely_integrable_periodic_offset) lemma absolutely_integrable_mult_Fejer_kernel_reflected2: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" shows "(\x. Fejer_kernel n x * f(t - x)) absolutely_integrable_on {-pi..pi}" proof - have "(\x. f(t - x)) absolutely_integrable_on {-pi..pi}" using assms apply (subst absolutely_integrable_reflect_real [symmetric]) apply (simp add: absolutely_integrable_periodic_offset) done then show ?thesis by (rule absolutely_integrable_mult_Fejer_kernel) qed lemma absolutely_integrable_mult_Fejer_kernel_reflected3: shows "(\x. Fejer_kernel n x * c) absolutely_integrable_on {-pi..pi}" using absolutely_integrable_on_const absolutely_integrable_mult_Fejer_kernel by blast lemma absolutely_integrable_mult_Fejer_kernel_reflected_part1: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" and "d \ pi" shows "(\x. Fejer_kernel n x * f(t + x)) absolutely_integrable_on {0..d}" by (rule absolutely_integrable_on_subinterval [OF absolutely_integrable_mult_Fejer_kernel_reflected1]) (auto simp: assms) lemma absolutely_integrable_mult_Fejer_kernel_reflected_part2: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" and "d \ pi" shows "(\x. Fejer_kernel n x * f(t - x)) absolutely_integrable_on {0..d}" by (rule absolutely_integrable_on_subinterval [OF absolutely_integrable_mult_Fejer_kernel_reflected2]) (auto simp: assms) lemma absolutely_integrable_mult_Fejer_kernel_reflected_part3: assumes "d \ pi" shows "(\x. Fejer_kernel n x * c) absolutely_integrable_on {0..d}" by (rule absolutely_integrable_on_subinterval [OF absolutely_integrable_mult_Fejer_kernel_reflected2]) (auto simp: assms) lemma absolutely_integrable_mult_Fejer_kernel_reflected_part4: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" and "d \ pi" shows "(\x. Fejer_kernel n x * (f(t + x) + f(t - x))) absolutely_integrable_on {0..d}" unfolding distrib_left by (intro set_integral_add absolutely_integrable_mult_Fejer_kernel_reflected_part1 absolutely_integrable_mult_Fejer_kernel_reflected_part2 assms) lemma absolutely_integrable_mult_Fejer_kernel_reflected_part5: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" and "d \ pi" shows "(\x. Fejer_kernel n x * ((f(t + x) + f(t - x)) - c)) absolutely_integrable_on {0..d}" unfolding distrib_left right_diff_distrib by (intro set_integral_add set_integral_diff absolutely_integrable_on_const absolutely_integrable_mult_Fejer_kernel_reflected_part1 absolutely_integrable_mult_Fejer_kernel_reflected_part2 assms, auto) lemma Fourier_sum_offset_Fejer_kernel_half: fixes n::nat assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" and "n > 0" shows "(\rk\2*r. Fourier_coefficient f k * trigonometric_set k t) / n - l = (LINT x|lebesgue_on {0..pi}. Fejer_kernel n x * (f(t + x) + f(t - x) - 2 * l)) / pi" proof - have "(\rk\2 * r. Fourier_coefficient f k * trigonometric_set k t) - real n * l = (\rk\2 * r. Fourier_coefficient f k * trigonometric_set k t) - l)" by (simp add: sum_subtractf) also have "\ = (\r = real n * ((LINT x|lebesgue_on {0..pi}. Fejer_kernel n x * (f(t + x) + f(t - x) - 2 * l)) / pi)" proof - have "integrable (lebesgue_on {0..pi}) (\x. Dirichlet_kernel i x * (f(t + x) + f(t - x) - 2 * l))" for i using absolutely_integrable_mult_Dirichlet_kernel_reflected_part2(2) f periodic by (force simp: intro!: absolutely_integrable_imp_integrable) then show ?thesis using \n > 0\ apply (simp add: Fejer_kernel_def flip: sum_divide_distrib) apply (simp add: sum_distrib_right flip: Bochner_Integration.integral_sum [symmetric]) done qed finally have "(\rk\2 * r. Fourier_coefficient f k * trigonometric_set k t) - real n * l = real n * ((LINT x|lebesgue_on {0..pi}. Fejer_kernel n x * (f(t + x) + f(t - x) - 2 * l)) / pi)" . with \n > 0\ show ?thesis by (auto simp: mult.commute divide_simps) qed lemma Fourier_sum_limit_Fejer_kernel_half: fixes n::nat assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" shows "(\n. ((\rk\2*r. Fourier_coefficient f k * trigonometric_set k t)) / n) \ l \ ((\n. integral\<^sup>L (lebesgue_on {0..pi}) (\x. Fejer_kernel n x * ((f(t + x) + f(t - x)) - 2*l))) \ 0)" (is "?lhs = ?rhs") proof - have "?lhs \ (\n. ((\rk\2*r. Fourier_coefficient f k * trigonometric_set k t)) / n - l) \ 0" by (simp add: LIM_zero_iff) also have "\ \ (\n. ((((\rk\2*r. Fourier_coefficient f k * trigonometric_set k t)) / n) - l) * pi) \ 0" using tendsto_mult_right_iff [OF pi_neq_zero] by simp also have "\ \ ?rhs" apply (intro Lim_transform_eq [OF Lim_transform_eventually [of "\n. 0"]] eventually_sequentiallyI [of 1]) apply (simp_all add: Fourier_sum_offset_Fejer_kernel_half assms) done finally show ?thesis . qed lemma has_integral_Fejer_kernel: "has_bochner_integral (lebesgue_on {-pi..pi}) (Fejer_kernel n) (if n = 0 then 0 else pi)" proof - have "has_bochner_integral (lebesgue_on {-pi..pi}) (\x. (\rrx. (\rr 0" by (simp add: Fejer_kernel) theorem Fourier_Fejer_Cesaro_summable: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" and fl: "(f \ l) (at t within atMost t)" and fr: "(f \ r) (at t within atLeast t)" shows "(\n. (\mk\2*m. Fourier_coefficient f k * trigonometric_set k t) / n) \ (l+r) / 2" proof - define h where "h \ \u. (f(t+u) + f(t-u)) - (l+r)" have "(\n. LINT u|lebesgue_on {0..pi}. Fejer_kernel n u * h u) \ 0" proof - have h0: "(h \ 0) (at 0 within atLeast 0)" proof - have l0: "((\u. f(t-u) - l) \ 0) (at 0 within {0..})" using fl unfolding Lim_within apply (elim all_forward imp_forward ex_forward conj_forward asm_rl, clarify) apply (drule_tac x="t-x" in bspec) apply (auto simp: dist_norm) done have r0: "((\u. f(t + u) - r) \ 0) (at 0 within {0..})" using fr unfolding Lim_within apply (elim all_forward imp_forward ex_forward conj_forward asm_rl, clarify) apply (drule_tac x="t+x" in bspec) apply (auto simp: dist_norm) done show ?thesis using tendsto_add [OF l0 r0] by (simp add: h_def algebra_simps) qed show ?thesis unfolding lim_sequentially dist_real_def diff_0_right proof clarify fix e::real assume "e > 0" then obtain x' where "0 < x'" "\x. \0 < x; x < x'\ \ \h x\ < e / (2 * pi)" using h0 unfolding Lim_within dist_real_def by (auto simp: dest: spec [where x="e/2/pi"]) then obtain \ where "0 < \" "\ < pi" and \: "\x. 0 < x \ x \ \ \ \h x\ < e/2/pi" apply (intro that [where \="min x' pi/2"], auto) using m2pi_less_pi by linarith have ftx: "(\x. f(t+x)) absolutely_integrable_on {-pi..pi}" using absolutely_integrable_periodic_offset assms by auto then have ftmx: "(\x. f(t-x)) absolutely_integrable_on {-pi..pi}" by (simp flip: absolutely_integrable_reflect_real [where f= "(\x. f(t+x))"]) have h_aint: "h absolutely_integrable_on {-pi..pi}" unfolding h_def by (intro absolutely_integrable_on_const set_integral_diff set_integral_add, auto simp: ftx ftmx) have "(\n. LINT x|lebesgue_on {\..pi}. Fejer_kernel n x * h x) \ 0" proof (rule Lim_null_comparison) define \ where "\ \ \n. (LINT x|lebesgue_on {\..pi}. \h x\ / (2 * sin(x/2) ^ 2)) / n" show "\\<^sub>F n in sequentially. norm (LINT x|lebesgue_on {\..pi}. Fejer_kernel n x * h x) \ \ n" proof (rule eventually_sequentiallyI) fix n::nat assume "n \ 1" have hint: "(\x. h x / (2 * sin(x/2) ^ 2)) absolutely_integrable_on {\..pi}" unfolding divide_inverse mult.commute [of "h _"] proof (rule absolutely_integrable_bounded_measurable_product_real) have cont: "continuous_on {\..pi} (\x. inverse (2 * (sin (x * inverse 2))\<^sup>2))" using \0 < \\ by (intro continuous_intros) (auto simp: sin_zero_pi_iff) show "(\x. inverse (2 * (sin (x * inverse 2))\<^sup>2)) \ borel_measurable (lebesgue_on {\..pi})" using \0 < \\ by (intro cont continuous_imp_measurable_on_sets_lebesgue) auto show "bounded ((\x. inverse (2 * (sin (x * inverse 2))\<^sup>2)) ` {\..pi})" using cont by (simp add: compact_continuous_image compact_imp_bounded) show "h absolutely_integrable_on {\..pi}" using \0 < \\ \\ < pi\ by (auto intro: absolutely_integrable_on_subinterval [OF h_aint]) qed auto then have *: "integrable (lebesgue_on {\..pi}) (\x. \h x\ / (2 * (sin (x/2))\<^sup>2))" by (simp add: absolutely_integrable_measurable o_def) define \ where "\ \ \x. (if n = 0 then 0 else if x = 0 then n/2 else (sin (real n / 2 * x))\<^sup>2 / (2 * real n * (sin (x/2))\<^sup>2)) * h x" have "\LINT x|lebesgue_on {\..pi}. \ x\ \ (LINT x|lebesgue_on {\..pi}. \h x\ / (2 * (sin (x/2))\<^sup>2) / n)" proof (rule integral_abs_bound_integral) show **: "integrable (lebesgue_on {\..pi}) (\x. \h x\ / (2 * (sin (x/2))\<^sup>2) / n)" using Bochner_Integration.integrable_mult_left [OF *, of "1/n"] by (simp add: field_simps) show \: "\\ x\ \ \h x\ / (2 * (sin (x/2))\<^sup>2) / real n" if "x \ space (lebesgue_on {\..pi})" for x using that \0 < \\ apply (simp add: \_def divide_simps mult_less_0_iff abs_mult) apply (auto simp: square_le_1 mult_left_le_one_le) done show "integrable (lebesgue_on {\..pi}) \" proof (rule measurable_bounded_by_integrable_imp_lebesgue_integrable [OF _ **]) let ?g = "\x. \h x\ / (2 * sin(x/2) ^ 2) / n" have ***: "integrable (lebesgue_on {\..pi}) (\x. (sin (n/2 * x))\<^sup>2 * (inverse (2 * (sin (x/2))\<^sup>2) * h x))" proof (rule absolutely_integrable_imp_integrable [OF absolutely_integrable_bounded_measurable_product_real]) show "(\x. (sin (real n / 2 * x))\<^sup>2) \ borel_measurable (lebesgue_on {\..pi})" by (intro continuous_imp_measurable_on_sets_lebesgue continuous_intros) auto show "bounded ((\x. (sin (real n / 2 * x))\<^sup>2) ` {\..pi})" by (force simp: square_le_1 intro: boundedI [where B=1]) show "(\x. inverse (2 * (sin (x/2))\<^sup>2) * h x) absolutely_integrable_on {\..pi}" using hint by (simp add: divide_simps) qed auto show "\ \ borel_measurable (lebesgue_on {\..pi})" apply (rule borel_measurable_integrable) apply (rule Bochner_Integration.integrable_cong [where f = "\x. sin(n / 2 * x) ^ 2 / (2 * n * sin(x/2) ^ 2) * h x", OF refl, THEN iffD1]) using \0 < \\ ** apply (force simp: \_def divide_simps algebra_simps mult_less_0_iff abs_mult) using Bochner_Integration.integrable_mult_left [OF ***, of "1/n"] by (simp add: field_simps) show "norm (\ x) \ ?g x" if "x \ {\..pi}" for x using that \ by (simp add: \_def) qed auto qed then show "norm (LINT x|lebesgue_on {\..pi}. Fejer_kernel n x * h x) \ \ n" by (simp add: Fejer_kernel \_def \_def flip: Bochner_Integration.integral_divide [OF *]) qed show "\ \ 0" unfolding \_def divide_inverse by (simp add: tendsto_mult_right_zero lim_inverse_n) qed then obtain N where N: "\n. n \ N \ \LINT x|lebesgue_on {\..pi}. Fejer_kernel n x * h x\ < e/2" unfolding lim_sequentially by (metis half_gt_zero_iff norm_conv_dist real_norm_def \e > 0\) show "\N. \n\N. \(LINT u|lebesgue_on {0..pi}. Fejer_kernel n u * h u)\ < e" proof (intro exI allI impI) fix n :: nat assume n: "n \ max 1 N" with N have 1: "\LINT x|lebesgue_on {\..pi}. Fejer_kernel n x * h x\ < e/2" by simp have "integral\<^sup>L (lebesgue_on {0..\}) (Fejer_kernel n) \ integral\<^sup>L (lebesgue_on {0..pi}) (Fejer_kernel n)" using \\ < pi\ has_bochner_integral_iff has_integral_Fejer_kernel_half by (force intro!: integral_mono_lebesgue_on_AE) also have "\ \ pi" using has_integral_Fejer_kernel_half by (simp add: has_bochner_integral_iff) finally have int_le_pi: "integral\<^sup>L (lebesgue_on {0..\}) (Fejer_kernel n) \ pi" . have 2: "\LINT x|lebesgue_on {0..\}. Fejer_kernel n x * h x\ \ (LINT x|lebesgue_on {0..\}. Fejer_kernel n x * e/2/pi)" proof - have eq: "integral\<^sup>L (lebesgue_on {0..\}) (\x. Fejer_kernel n x * h x) = integral\<^sup>L (lebesgue_on {0..\}) (\x. Fejer_kernel n x * (if x = 0 then 0 else h x))" proof (rule integral_cong_AE) have \: "{u. u \ 0 \ P u} \ {0..\} = {0} \ Collect P \ {0..\}" "{u. u \ 0 \ P u} \ {0..\} = (Collect P \ {0..\}) - {0}" for P using \0 < \\ by auto have *: "- {0} \ A \ {0..\} = A \ {0..\} - {0}" for A by auto show "(\x. Fejer_kernel n x * h x) \ borel_measurable (lebesgue_on {0..\})" using \\ < pi\ by (intro absolutely_integrable_imp_borel_measurable h_aint absolutely_integrable_on_subinterval [OF absolutely_integrable_mult_Fejer_kernel], auto) then show "(\x. Fejer_kernel n x * (if x = 0 then 0 else h x)) \ borel_measurable (lebesgue_on {0..\})" apply (simp add: in_borel_measurable Ball_def vimage_def Collect_conj_eq Collect_imp_eq * flip: Collect_neg_eq) apply (elim all_forward imp_forward asm_rl) using \0 < \\ apply (auto simp: \ sets.insert_in_sets sets_restrict_space_iff cong: conj_cong) done have 0: "{0} \ null_sets (lebesgue_on {0..\})" using \0 < \\ by (simp add: null_sets_restrict_space) then show "AE x in lebesgue_on {0..\}. Fejer_kernel n x * h x = Fejer_kernel n x * (if x = 0 then 0 else h x)" by (auto intro: AE_I' [OF 0]) qed show ?thesis unfolding eq proof (rule integral_abs_bound_integral) have "(\x. if x = 0 then 0 else h x) absolutely_integrable_on {- pi..pi}" proof (rule absolutely_integrable_spike [OF h_aint]) show "negligible {0}" by auto qed auto with \0 < \\ \\ < pi\ show "integrable (lebesgue_on {0..\}) (\x. Fejer_kernel n x * (if x = 0 then 0 else h x))" by (intro absolutely_integrable_imp_integrable h_aint absolutely_integrable_on_subinterval [OF absolutely_integrable_mult_Fejer_kernel]) auto show "integrable (lebesgue_on {0..\}) (\x. Fejer_kernel n x * e / 2 / pi)" by (simp add: absolutely_integrable_imp_integrable \\ < pi\ absolutely_integrable_mult_Fejer_kernel_reflected_part3 less_eq_real_def) show "\Fejer_kernel n x * (if x = 0 then 0 else h x)\ \ Fejer_kernel n x * e / 2 / pi" if "x \ space (lebesgue_on {0..\})" for x using that \ [of x] \e > 0\ by (auto simp: abs_mult eq simp flip: times_divide_eq_right intro: mult_left_mono) qed qed have 3: "\ \ e/2" using int_le_pi \0 < e\ by (simp add: divide_simps mult.commute [of e]) have "integrable (lebesgue_on {0..pi}) (\x. Fejer_kernel n x * h x)" unfolding h_def by (simp add: absolutely_integrable_imp_integrable absolutely_integrable_mult_Fejer_kernel_reflected_part5 assms) then have "LINT x|lebesgue_on {0..pi}. Fejer_kernel n x * h x = (LINT x|lebesgue_on {0..\}. Fejer_kernel n x * h x) + (LINT x|lebesgue_on {\..pi}. Fejer_kernel n x * h x)" by (rule integral_combine) (use \0 < \\ \\ < pi\ in auto) then show "\LINT u|lebesgue_on {0..pi}. Fejer_kernel n u * h u\ < e" using 1 2 3 by linarith qed qed qed then show ?thesis unfolding h_def by (simp add: Fourier_sum_limit_Fejer_kernel_half assms add_divide_distrib) qed corollary Fourier_Fejer_Cesaro_summable_simple: assumes f: "continuous_on UNIV f" and periodic: "\x. f(x + 2*pi) = f x" shows "(\n. (\mk\2*m. Fourier_coefficient f k * trigonometric_set k x) / n) \ f x" proof - have "(\n. (\mk\2*m. Fourier_coefficient f k * trigonometric_set k x) / n) \ (f x + f x) / 2" proof (rule Fourier_Fejer_Cesaro_summable) show "f absolutely_integrable_on {- pi..pi}" using absolutely_integrable_continuous_real continuous_on_subset f by blast show "(f \ f x) (at x within {..x})" "(f \ f x) (at x within {x..})" using Lim_at_imp_Lim_at_within continuous_on_def f by blast+ qed (auto simp: periodic Lim_at_imp_Lim_at_within continuous_on_def f) then show ?thesis by simp qed end diff --git a/thys/Green/CircExample.thy b/thys/Green/CircExample.thy --- a/thys/Green/CircExample.thy +++ b/thys/Green/CircExample.thy @@ -1,794 +1,794 @@ section \The Circle Example\ theory CircExample imports Green SymmetricR2Shapes begin locale circle = R2 + fixes d::real assumes d_gt_0: "0 < d" begin definition circle_y where "circle_y t = sqrt (1/4 - t * t)" definition circle_cube where "circle_cube = (\(x,y). ((x - 1/2) * d, (2 * y - 1) * d * sqrt (1/4 - (x -1/2)*(x -1/2))))" lemma circle_cube_nice: shows "circle_cube = (\(x,y). (d * x_coord x, (2 * y - 1) * d * circle_y (x_coord x)))" by (auto simp add: circle_cube_def circle_y_def x_coord_def) definition rot_circle_cube where "rot_circle_cube = prod.swap \ (circle_cube) \ prod.swap" abbreviation "rot_y t1 t2 \ ((t1-1/2)/(2* circle_y (x_coord (rot_x t1 t2))) +1/2::real)" definition "x_coord_inv (x::real) = (1/2) + x" lemma x_coord_inv_1: "x_coord_inv (x_coord (x::real)) = x" by (auto simp add: x_coord_inv_def x_coord_def) lemma x_coord_inv_2: "x_coord (x_coord_inv (x::real)) = x" by (auto simp add: x_coord_inv_def x_coord_def) definition "circle_y_inv = circle_y" abbreviation "rot_x'' (x::real) (y::real) \ (x_coord_inv ((2 * y - 1) * circle_y (x_coord x)))" lemma circle_y_bounds: assumes "-1/2 \ (x::real) \ x \ 1/2" shows "0 \ circle_y x" "circle_y x \ 1/2" unfolding circle_y_def real_sqrt_ge_0_iff proof - show "0 \ 1/4 - x * x" - using assms by (sos "(((A<0 * R<1) + ((A<=0 * (A<=1 * R<1)) * (R<1 * [1]^2))))") + using assms + by (sos "(((A<0 * R<1) + ((A<=0 * (A<=1 * R<1)) * (R<1 * [1]^2))))") show "sqrt (1/4 - x * x) \ 1/2" apply (rule real_le_lsqrt) - using assms apply(auto simp add: divide_simps algebra_simps) - by (sos "(((A<0 * R<1) + ((A<=0 * (A<=1 * R<1)) * (R<1 * [1]^2))))") + using assms by(auto simp add: divide_simps algebra_simps) qed lemma circle_y_x_coord_bounds: assumes "0 \ (x::real)" "x \ 1" shows "0 \ circle_y (x_coord x) \ circle_y (x_coord x) \ 1/2" using circle_y_bounds[OF x_coord_bounds[OF assms]] by auto lemma rot_x_ivl: assumes "(0::real) \ x" "x \ 1" "0 \ y" "y \ 1" shows "0 \ rot_x'' x y \ rot_x'' x y \ 1" proof have "\a::real. 0 \ a \ a \ 1/2 \ 0 \ 1/2 + (2 * y - 1) * a" using assms by (sos "(((A<0 * R<1) + (((A<=4 * R<1) * (R<1/2 * [1]^2)) + (((A<=1 * (A<=5 * R<1)) * (R<1 * [1]^2)) + ((A<=0 * (A<=4 * R<1)) * (R<1 * [1]^2))))))") then show "0 \ rot_x'' x y" using assms circle_y_x_coord_bounds by(auto simp add: x_coord_inv_def) have "\a::real. 0 \ a \ a \ 1/2 \ 1/2 + (2 * y - 1) * a \ 1" using assms by (sos "(((A<0 * R<1) + (((A<=5 * R<1) * (R<1/2 * [1]^2)) + (((A<=1 * (A<=4 * R<1)) * (R<1 * [1]^2)) + ((A<=0 * (A<=5 * R<1)) * (R<1 * [1]^2))))))") then show "rot_x'' x y \ 1" using assms circle_y_x_coord_bounds by (auto simp add: x_coord_inv_def) qed abbreviation "rot_y'' (x::real) (y::real) \ (x_coord x)/(2* (circle_y (x_coord (rot_x'' x y)))) + 1/2" lemma rot_y_ivl: assumes "(0::real) \ x" "x \ 1" "0 \ y" "y \ 1" shows "0 \ rot_y'' x y \ rot_y'' x y \ 1" proof show "0 \ rot_y'' x y" proof(cases "(x_coord x) < 0") case True have i: "\a b::real. a < 0 \ 0 \ a + b \ (0 \ a/(2*(b)) + 1/2)" by(auto simp add: algebra_simps divide_simps) have *: "(1/2 - x) \ sqrt (x * x + (1/4 + (x * (y * 4) + x * (x * (y * (y * 4))))) - (x + (x * (x * (y * 4)) + x * (y * (y * 4)))))" apply (rule real_le_rsqrt) using assms apply (simp add: algebra_simps power2_eq_square mult_left_le_one_le) by (sos "(((A<0 * R<1) + ((A<=0 * (A<=1 * (A<=2 * (A<=3 * R<1)))) * (R<4 * [1]^2))))") have rw: "\x - x * x\ = x - x * x" using assms by (sos "(() & (((((A<0 * A<1) * R<1) + ((A<=0 * (A<=1 * (A<1 * R<1))) * (R<1 * [1]^2)))) & ((((A<0 * A<1) * R<1) + ((A<=0 * (A<=1 * (A<1 * R<1))) * (R<1 * [1]^2))))))") have "0 \ x_coord x + (circle_y (x_coord (rot_x'' x y)))" using * apply (auto simp add: x_coord_inv_2) by (auto simp add: circle_y_def algebra_simps rw x_coord_def) then show ?thesis using True i by blast next case False have i: "\a b::real. 0 \ a \ 0 \ b \ (0 \ a/(2*b) + 1/2)" by(auto simp add: algebra_simps divide_simps) have "0 \ circle_y (x_coord (x_coord_inv ((2 * y - 1) * circle_y (x_coord x))))" proof - have rw: "\x - x * x\ = x - x * x" using assms by (sos "(() & (((((A<0 * A<1) * R<1) + ((A<=0 * (A<=1 * (A<1 * R<1))) * (R<1 * [1]^2)))) & ((((A<0 * A<1) * R<1) + ((A<=0 * (A<=1 * (A<1 * R<1))) * (R<1 * [1]^2))))))") have "\x. 0 \ x \ x \ 1/2 \ -1/2\ (2 * y - 1) * x" using assms by (sos "(((A<0 * R<1) + (((A<=4 * R<1) * (R<1/2 * [1]^2)) + (((A<=1 * (A<=5 * R<1)) * (R<1 * [1]^2)) + ((A<=0 * (A<=4 * R<1)) * (R<1 * [1]^2))))))") then have "- 1/2 \ (2 * y - 1) * circle_y (x_coord x)" using circle_y_x_coord_bounds assms(1-2) by auto moreover have "\x. 0 \ x \ x \ 1/2 \ (2 * y - 1) * x \ 1/2 " using assms by (sos "(((A<0 * R<1) + (((A<=5 * R<1) * (R<1/2 * [1]^2)) + (((A<=1 * (A<=4 * R<1)) * (R<1 * [1]^2)) + ((A<=0 * (A<=5 * R<1)) * (R<1 * [1]^2))))))") then have "(2 * y - 1) * circle_y (x_coord x) \ 1/2" using circle_y_x_coord_bounds assms(1-2) by auto ultimately show "0 \ circle_y (x_coord (x_coord_inv ((2 * y - 1) * circle_y (x_coord x))))" by (simp add: circle_y_bounds(1) x_coord_inv_2) qed then show ?thesis using False by auto qed have i: "\a b::real. a < 0 \ 0 \ b \ (a/(2*(b)) + 1/2) \ 1" by(auto simp add: algebra_simps divide_simps) show "rot_y'' x y \ 1" proof(cases "(x_coord x) < 0") case True have i: "\a b::real. a < 0 \ 0 \ b \ (a/(2*(b)) + 1/2) \ 1" by(auto simp add: algebra_simps divide_simps) have "\x. 0 \ x \ x \ 1/2 \ -1/2\ (2 * y - 1) * x" using assms by (sos "(((A<0 * R<1) + (((A<=4 * R<1) * (R<1/2 * [1]^2)) + (((A<=1 * (A<=5 * R<1)) * (R<1 * [1]^2)) + ((A<=0 * (A<=4 * R<1)) * (R<1 * [1]^2))))))") then have "- 1/2 \ (2 * y - 1) * circle_y (x_coord x)" using circle_y_x_coord_bounds assms(1-2) by auto moreover have "\x. 0 \ x \ x \ 1/2 \ (2 * y - 1) * x \ 1/2 " using assms by (sos "(((A<0 * R<1) + (((A<=5 * R<1) * (R<1/2 * [1]^2)) + (((A<=1 * (A<=4 * R<1)) * (R<1 * [1]^2)) + ((A<=0 * (A<=5 * R<1)) * (R<1 * [1]^2))))))") then have "(2 * y - 1) * circle_y (x_coord x) \ 1/2" using circle_y_x_coord_bounds assms(1-2) by auto ultimately have "0 \ circle_y (x_coord (x_coord_inv ((2 * y - 1) * circle_y (x_coord x))))" by (simp add: circle_y_bounds(1) x_coord_inv_2) then show ?thesis by (simp add: True i) next case False have i: "\a b::real. 0 \ a \ a \ b \ (a/(2*b) + 1/2) \ 1" by(auto simp add: algebra_simps divide_simps) have "(x - 1/2) * (x - 1/2) \ (x * x + (1/4 + (x * (y * 4) + x * (x * (y * (y * 4))))) - (x + (x * (x * (y * 4)) + x * (y * (y * 4)))))" using assms False apply(auto simp add: x_coord_def) by (sos "(((A<0 * R<1) + (((A<=0 * (A<=1 * (A<=2 * R<1))) * (R<2 * [1]^2)) + ((A<=0 * (A<=1 * (A<=2 * (A<=3 * R<1)))) * (R<2 * [1]^2)))))") then have "sqrt ((x - 1/2) * (x - 1/2)) \ sqrt (x * x + (1/4 + (x * (y * 4) + x * (x * (y * (y * 4))))) - (x + (x * (x * (y * 4)) + x * (y * (y * 4)))))" using real_sqrt_le_mono by blast then have *: "(x - 1/2) \ sqrt (x * x + (1/4 + (x * (y * 4) + x * (x * (y * (y * 4))))) - (x + (x * (x * (y * 4)) + x * (y * (y * 4)))))" using assms False by(auto simp add: x_coord_def) have rw: "\x - x * x\ = x - x * x" using assms by (sos "(() & (((((A<0 * A<1) * R<1) + ((A<=0 * (A<=1 * (A<1 * R<1))) * (R<1 * [1]^2)))) & ((((A<0 * A<1) * R<1) + ((A<=0 * (A<=1 * (A<1 * R<1))) * (R<1 * [1]^2))))))") have "x_coord x \ circle_y (x_coord (x_coord_inv ((2 * y - 1) * circle_y (x_coord x))))" using * unfolding x_coord_inv_2 by (auto simp add: circle_y_def algebra_simps rw x_coord_def) then show ?thesis using False i by auto qed qed lemma circle_eq_rot_circle: assumes "0 \ x" "x \ 1" "0 \ y" "y \ 1" shows "(circle_cube (x, y)) = (rot_circle_cube (rot_y'' x y, rot_x'' x y))" proof have rw: "\1/4 - x_coord x * x_coord x\ = 1/4 - x_coord x * x_coord x" apply(rule abs_of_nonneg) using assms mult_left_le by (auto simp add: x_coord_def divide_simps algebra_simps) show "fst (circle_cube (x, y)) = fst (rot_circle_cube (rot_y'' x y, rot_x'' x y))" using assms d_gt_0 apply(simp add: circle_cube_nice rot_circle_cube_def x_coord_inv_2 circle_y_def algebra_simps rw) apply(auto simp add: x_coord_def algebra_simps) by (sos "(((((A<0 * A<1) * ((A<0 * A<1) * R<1)) + (([~4*d^2] * A=0) + (((A<=1 * (A<=2 * (A<=3 * R<1))) * (R<8 * [d]^2)) + ((A<=1 * (A<=2 * (A<=3 * (A<1 * R<1)))) * (R<8 * [d]^2)))))) & ((((A<0 * A<1) * ((A<0 * A<1) * R<1)) + (([~4*d^2] * A=0) + (((A<=0 * (A<=2 * (A<=3 * R<1))) * (R<8 * [d]^2)) + ((A<=0 * (A<=2 * (A<=3 * (A<1 * R<1)))) * (R<8 * [d]^2)))))))") (*Be patient: it takes a bit of time, still better than creeping its proof manually.*) show "snd (circle_cube (x, y)) = snd (rot_circle_cube (rot_y'' x y, rot_x'' x y))" using assms by(auto simp add: circle_cube_def rot_circle_cube_def x_coord_inv_def circle_y_def x_coord_def) qed lemma rot_circle_eq_circle: assumes "0 \ x" "x \ 1" "0 \ y" "y \ 1" shows "(rot_circle_cube (x, y)) = (circle_cube (rot_x'' y x, rot_y'' y x))" proof show "fst (rot_circle_cube (x, y)) = fst (circle_cube (rot_x'' y x, rot_y'' y x))" using assms by(auto simp add: circle_cube_def rot_circle_cube_def x_coord_inv_def circle_y_def x_coord_def) have rw: "\1/4 - x_coord y * x_coord y\ = 1/4 - x_coord y * x_coord y" apply(rule abs_of_nonneg) using assms mult_left_le by (auto simp add: x_coord_def divide_simps algebra_simps) show "snd (rot_circle_cube (x, y)) = snd (circle_cube (rot_x'' y x, rot_y'' y x))" using assms d_gt_0 apply(simp add: circle_cube_nice rot_circle_cube_def x_coord_inv_2 circle_y_def algebra_simps rw) apply(auto simp add: x_coord_def algebra_simps) by (sos "(((((A<0 * A<1) * ((A<0 * A<1) * R<1)) + (([~4*d^2] * A=0) + (((A<=0 * (A<=1 * (A<=3 * R<1))) * (R<8 * [d]^2)) + ((A<=0 * (A<=1 * (A<=3 * (A<1 * R<1)))) * (R<8 * [d]^2)))))) & ((((A<0 * A<1) * ((A<0 * A<1) * R<1)) + (([~4*d^2] * A=0) + (((A<=0 * (A<=1 * (A<=2 * R<1))) * (R<8 * [d]^2)) + ((A<=0 * (A<=1 * (A<=2 * (A<1 * R<1)))) * (R<8 * [d]^2)))))))") (*Be patient: it takes a bit of time, still better than creeping its proof manually.*) qed lemma rot_img_eq: assumes "0 < d" shows "(cubeImage (circle_cube )) = (cubeImage (rot_circle_cube))" apply(auto simp add: cubeImage_def image_def cbox_def real_pair_basis) by (meson rot_y_ivl rot_x_ivl assms circle_eq_rot_circle rot_circle_eq_circle)+ lemma rot_circle_div_circle: assumes "0 < (d::real)" shows "gen_division (cubeImage circle_cube) (cubeImage ` {rot_circle_cube})" using rot_img_eq[OF assms] by(auto simp add: gen_division_def) lemma circle_cube_boundary_valid: assumes "(k,\)\boundary circle_cube" shows "valid_path \" proof - have f01: "finite{0,1}" by simp show ?thesis using assms unfolding boundary_def horizontal_boundary_def vertical_boundary_def circle_cube_def valid_path_def piecewise_C1_differentiable_on_def by safe (rule derivative_intros continuous_intros f01 exI ballI conjI refl | force simp add: field_simps)+ qed lemma rot_circle_cube_boundary_valid: assumes "(k,\)\boundary rot_circle_cube" shows "valid_path \" using assms swap_valid_boundaries circle_cube_boundary_valid by (fastforce simp add: rot_circle_cube_def) lemma diff_divide_cancel: fixes z::real shows "z \ 0 \ (a * z - a * (b * z)) / z = (a - a * b)" by (auto simp: field_simps) lemma circle_cube_is_type_I: assumes "0 < d" shows "typeI_twoCube circle_cube" unfolding typeI_twoCube_def proof (intro exI conjI ballI) have f01: "finite{-d/2,d/2}" by simp show "-d/2 < d/2" using assms by simp show "(\x. d * sqrt (1/4 - (x/d) * (x/d))) piecewise_C1_differentiable_on {- d/2..d/2}" using assms unfolding piecewise_C1_differentiable_on_def apply (intro exI conjI) apply (rule ballI refl f01 derivative_intros continuous_intros | simp)+ apply (auto simp: field_simps) by sos show "(\x. -d * sqrt (1/4 - (x/d) * (x/d))) piecewise_C1_differentiable_on {- d/2..d/2}" using assms unfolding piecewise_C1_differentiable_on_def apply (intro exI conjI) apply (rule ballI refl f01 derivative_intros continuous_intros | simp)+ apply (auto simp: field_simps) by sos show "- d * sqrt (1/4 - x / d * (x / d)) \ d * sqrt (1/4 - x / d * (x / d))" if "x \ {- d/2..d/2}" for x proof - have *: "x^2 \ (d/2)^2" using real_sqrt_le_iff that by fastforce show ?thesis apply (rule mult_right_mono) using assms * apply (simp_all add: divide_simps power2_eq_square) done qed qed (auto simp add: circle_cube_def divide_simps algebra_simps diff_divide_cancel) lemma rot_circle_cube_is_type_II: shows "typeII_twoCube rot_circle_cube" using d_gt_0 swap_typeI_is_typeII circle_cube_is_type_I by (auto simp add: rot_circle_cube_def) definition circle_bot_edge where "circle_bot_edge = (1::int, \t. (x_coord t * d, - d * circle_y (x_coord t)))" definition circle_top_edge where "circle_top_edge = (- 1::int, \t. (x_coord t * d, d * circle_y (x_coord t)))" definition circle_right_edge where "circle_right_edge = (1::int, \y. (d/2, 0))" definition circle_left_edge where "circle_left_edge = (- 1::int, \y. (- (d/2), 0))" lemma circle_cube_boundary_explicit: "boundary circle_cube = {circle_left_edge,circle_right_edge,circle_bot_edge,circle_top_edge}" by (auto simp add: valid_two_cube_def boundary_def horizontal_boundary_def vertical_boundary_def circle_cube_def circle_top_edge_def circle_bot_edge_def circle_cube_nice x_coord_def circle_y_def circle_left_edge_def circle_right_edge_def) definition rot_circle_right_edge where "rot_circle_right_edge = (1::int, \t. (d * circle_y (x_coord t), x_coord t * d))" definition rot_circle_left_edge where "rot_circle_left_edge = (- 1::int, \t. (- d * circle_y (x_coord t), x_coord t * d))" definition rot_circle_top_edge where "rot_circle_top_edge = (- 1::int, \y. (0, d/2))" definition rot_circle_bot_edge where "rot_circle_bot_edge = (1::int, \y. (0, - (d/2)))" lemma rot_circle_cube_boundary_explicit: "boundary (rot_circle_cube) = {rot_circle_top_edge,rot_circle_bot_edge,rot_circle_right_edge,rot_circle_left_edge}" by (auto simp add: rot_circle_cube_def valid_two_cube_def boundary_def horizontal_boundary_def vertical_boundary_def circle_cube_def rot_circle_right_edge_def rot_circle_left_edge_def x_coord_def circle_y_def rot_circle_top_edge_def rot_circle_bot_edge_def) lemma rot_circle_cube_vertical_boundary_explicit: "vertical_boundary rot_circle_cube = {rot_circle_right_edge,rot_circle_left_edge}" by (auto simp add: rot_circle_cube_def valid_two_cube_def vertical_boundary_def circle_cube_def rot_circle_right_edge_def rot_circle_left_edge_def x_coord_def circle_y_def) lemma circ_left_edge_neq_top: "(- 1::int, \y::real. (- (d/2), 0)) \ (- 1, \x. ((x - 1/2) * d, d * sqrt (1/4 - (x - 1/2) * (x - 1/2))))" by (metis (no_types, lifting) add_diff_cancel_right' d_gt_0 mult.commute mult_cancel_left order_less_irrefl prod.inject) lemma circle_cube_valid_two_cube: "valid_two_cube (circle_cube)" proof (auto simp add: valid_two_cube_def boundary_def horizontal_boundary_def vertical_boundary_def circle_cube_def) have iv: "(- 1::int, \y::real. (- (d/2), 0)) \ (- 1, \x. ((x - 1/2) * d, d * sqrt (1/4 - (x - 1/2) * (x - 1/2))))" using d_gt_0 apply (auto simp add: algebra_simps) by (metis (no_types, opaque_lifting) add_diff_cancel_right' add_uminus_conv_diff cancel_comm_monoid_add_class.diff_cancel less_eq_real_def linorder_not_le mult.left_neutral prod.simps(1)) have v: "(1::int, \y. (d/2, 0)) \ (1, \x. ((x - 1/2) * d, - (d * sqrt (1/4 - (x - 1/2) * (x - 1/2)))))" using d_gt_0 apply (auto simp add: algebra_simps) by (metis (no_types, opaque_lifting) diff_0 equal_neg_zero mult_zero_left nonzero_mult_div_cancel_left order_less_irrefl prod.sel(1) times_divide_eq_right zero_neq_numeral) show " card {(- 1::int, \y. (- (d/2), 0)), (1, \y. (d/2, 0)), (1, \x. ((x - 1/2) * d, - (d * sqrt (1/4 - (x - 1/2) * (x - 1/2))))), (- 1, \x. ((x - 1/2) * d, d * sqrt (1/4 - (x - 1/2) * (x - 1/2))))} = 4" using iv v by auto qed lemma rot_circle_cube_valid_two_cube: shows "valid_two_cube rot_circle_cube" using valid_cube_valid_swap circle_cube_valid_two_cube d_gt_0 rot_circle_cube_def by force definition circle_arc_0 where "circle_arc_0 = (1, \t::real. (0,0))" lemma circle_top_bot_edges_neq' [simp]: shows "circle_top_edge \ circle_bot_edge" by (simp add: circle_top_edge_def circle_bot_edge_def) lemma rot_circle_top_left_edges_neq [simp]: "rot_circle_top_edge \ rot_circle_left_edge" apply (simp add: rot_circle_left_edge_def rot_circle_top_edge_def x_coord_def) by (metis (mono_tags, opaque_lifting) cancel_comm_monoid_add_class.diff_cancel d_gt_0 divide_eq_0_iff mult_zero_left order_less_irrefl prod.sel(2) zero_neq_numeral) lemma rot_circle_bot_left_edges_neq [simp]: "rot_circle_bot_edge \ rot_circle_left_edge" by (simp add: rot_circle_left_edge_def rot_circle_bot_edge_def x_coord_def) lemma rot_circle_top_right_edges_neq [simp]: "rot_circle_top_edge \ rot_circle_right_edge" by (simp add: rot_circle_right_edge_def rot_circle_top_edge_def x_coord_def) lemma rot_circle_bot_right_edges_neq [simp]: "rot_circle_bot_edge \ rot_circle_right_edge" apply (simp add: rot_circle_right_edge_def rot_circle_bot_edge_def x_coord_def) by (metis (mono_tags, opaque_lifting) cancel_comm_monoid_add_class.diff_cancel d_gt_0 divide_eq_0_iff mult_zero_left neg_0_equal_iff_equal order_less_irrefl prod.sel(2) zero_neq_numeral) lemma rot_circle_right_top_edges_neq' [simp]: "rot_circle_right_edge \ rot_circle_left_edge" by (simp add: rot_circle_left_edge_def rot_circle_right_edge_def) lemma rot_circle_left_bot_edges_neq [simp]: "rot_circle_left_edge \ rot_circle_top_edge" apply (simp add: rot_circle_top_edge_def rot_circle_left_edge_def) by (metis (no_types, opaque_lifting) cancel_comm_monoid_add_class.diff_cancel d_gt_0 mult.commute mult_zero_right nonzero_mult_div_cancel_left order_less_irrefl prod.sel(2) times_divide_eq_right x_coord_def zero_neq_numeral) lemma circle_right_top_edges_neq [simp]: "circle_right_edge \ circle_top_edge" proof - have "circle_right_edge = (1, (\r::real. (d / 2, 0::real)))" using circle.circle_right_edge_def circle_axioms by blast then show ?thesis using circle.circle_top_edge_def circle_axioms by auto qed lemma circle_left_bot_edges_neq [simp]: "circle_left_edge \ circle_bot_edge" proof - have "circle_bot_edge = (1, \r. (x_coord r * d, - d * circle_y (x_coord r)))" using circle.circle_bot_edge_def circle_axioms by blast then show ?thesis by (simp add: circle_left_edge_def) qed lemma circle_left_top_edges_neq [simp]: "circle_left_edge \ circle_top_edge" proof - have "\r. ((r - 1 / 2) * d, d * sqrt (1 / 4 - (r - 1 / 2) * (r - 1 / 2))) \ (- (d / 2), 0)" by (metis circ_left_edge_neq_top) then have "(\r. d * r \ - (d / 2)) \ (\r ra. (x_coord (x_coord_inv r) * d, d * circle_y (x_coord (x_coord_inv r))) = (x_coord ra * d, d * circle_y (x_coord ra)) \ d * circle_y r \ 0)" by (metis mult.commute) then have "(\r. (x_coord r * d, d * circle_y (x_coord r))) \ (\r. (- (d / 2), 0))" by (metis mult.commute prod.simps(1) x_coord_inv_2) then show ?thesis by (simp add: circle_left_edge_def circle_top_edge_def) qed lemma circle_right_bot_edges_neq [simp]: "circle_right_edge \ circle_bot_edge" by (smt Pair_inject circle_bot_edge_def d_gt_0 circle.circle_right_edge_def circle_axioms mult_le_cancel_iff1 x_coord_def) definition circle_polar where "circle_polar t = ((d/2) * cos (2 * pi * t), (d/2) * sin (2 * pi * t))" lemma circle_polar_smooth: "(circle_polar) C1_differentiable_on {0..1}" proof - have "inj ((*) (2 * pi))" by (auto simp: inj_on_def) then have *: "\x. finite ({0..1} \ (*) (2 * pi) -` {x})" by (simp add: finite_vimageI) have "(\t. ((d/2) * cos (2 * pi * t), (d/2) * sin (2 * pi * t))) C1_differentiable_on {0..1}" by (rule * derivative_intros)+ then show ?thesis apply (rule eq_smooth_gen) by(simp add: circle_polar_def) qed abbreviation "custom_arccos \ (\x. (if -1 \ x \ x \ 1 then arccos x else (if x < -1 then -x + pi else 1 -x )))" lemma cont_custom_arccos: assumes "S \ {-1..1}" shows "continuous_on S custom_arccos" proof - have "continuous_on ({-1..1} \ {}) custom_arccos" by (auto intro!: continuous_on_cases continuous_intros) with assms show ?thesis using continuous_on_subset by auto qed lemma custom_arccos_has_deriv: assumes "- 1 < x" "x < 1" shows "DERIV custom_arccos x :> inverse (- sqrt (1 - x\<^sup>2))" proof - have x1: "\x\\<^sup>2 < 1\<^sup>2" by (simp add: abs_less_iff abs_square_less_1 assms) show ?thesis apply (rule DERIV_inverse_function [where f=cos and a="-1" and b=1]) apply (rule x1 derivative_eq_intros | simp add: sin_arccos)+ using assms x1 cont_custom_arccos [of "{-1<..<1}"] apply (auto simp: continuous_on_eq_continuous_at greaterThanLessThan_subseteq_atLeastAtMost_iff) done qed declare custom_arccos_has_deriv[THEN DERIV_chain2, derivative_intros] custom_arccos_has_deriv[THEN DERIV_chain2, unfolded has_field_derivative_def, derivative_intros] lemma circle_boundary_reparams: shows rot_circ_left_edge_reparam_polar_circ_split: "reparam (rec_join [(rot_circle_left_edge)]) (rec_join [(subcube (1/4) (1/2) (1, circle_polar)), (subcube (1/2) (3/4) (1, circle_polar))])" (is ?P1) and circ_top_edge_reparam_polar_circ_split: "reparam (rec_join [(circle_top_edge)]) (rec_join [(subcube 0 (1/4) (1, circle_polar)), (subcube (1/4) (1/2) (1, circle_polar))])" (is ?P2) and circ_bot_edge_reparam_polar_circ_split: "reparam (rec_join [(circle_bot_edge)]) (rec_join [(subcube (1/2) (3/4) (1, circle_polar)), (subcube (3/4) 1 (1, circle_polar))])" (is ?P3) and rot_circ_right_edge_reparam_polar_circ_split: "reparam (rec_join [(rot_circle_right_edge)]) (rec_join [(subcube (3/4) 1 (1, circle_polar)), (subcube 0 (1/4) (1, circle_polar))])" (is ?P4) proof - let ?\ = "((*) (1/pi) \ custom_arccos \ (\t. 2 * x_coord (1 - t)))" let ?LHS1 = "(\x. (- (d * sqrt (1/4 - x_coord (1 - x) * x_coord (1 - x))), x_coord (1 - x) * d))" let ?RHS1 = "((\x. if x * 2 \ 1 then (d * cos (2 * pi * (2 * x/4 + 1/4))/2, d * sin (2 * pi * (2 * x/4 + 1/4))/2) else (d * cos (2 * pi * ((2 * x - 1)/4 + 1/2))/2, d * sin (2 * pi * ((2 * x - 1)/4 + 1/2))/2)) \ ?\)" let ?LHS2 = "\x. ((x_coord (1 - x) * d, d * sqrt (1/4 - x_coord (1 - x) * x_coord (1 - x))))" let ?RHS2 = "((\x. if x * 2 \ 1 then (d * cos (2 * x * pi/2)/2, d * sin (2 * x * pi/2)/2) else (d * cos (2 * pi * ((2 * x - 1)/4 + 1/4))/2, d * sin (2 * pi * ((2 * x - 1)/4 + 1/4))/2)) \ ?\)" let ?LHS3 = "\x. (x_coord x * d, - (d * sqrt (1/4 - x_coord x * x_coord x)))" let ?RHS3 = "(\x. if x * 2 \ 1 then (d * cos (2 * pi * (2 * x/4 + 1/2))/2, d * sin (2 * pi * (2 * x/4 + 1/2))/2) else (d * cos (2 * pi * ((2 * x - 1)/4 + 3/4))/2, d * sin (2 * pi * ((2 * x - 1)/4 + 3/4))/2))" let ?LHS4 = "\x. (d * sqrt (1/4 - x_coord x * x_coord x), x_coord x * d)" let ?RHS4 = "(\x. if x * 2 \ 1 then (d * cos (2 * pi * (2 * x/4 + 3/4))/2, d * sin (2 * pi * (2 * x/4 + 3/4))/2) else (d * cos ((2 * x - 1) * pi/2)/2, d * sin ((2 * x - 1) * pi/2)/2))" have phi_diff: "?\ piecewise_C1_differentiable_on {0..1}" unfolding piecewise_C1_differentiable_on_def proof show "continuous_on {0..1} ?\" unfolding x_coord_def by (intro continuous_on_compose cont_custom_arccos continuous_intros) auto have "?\ C1_differentiable_on {0..1} - {0,1}" unfolding x_coord_def piecewise_C1_differentiable_on_def C1_differentiable_on_def valid_path_def by (simp | rule has_vector_derivative_pair_within DERIV_image_chain derivative_eq_intros continuous_intros exI ballI conjI | force simp add: field_simps | sos)+ then show "\s. finite s \ ?\ C1_differentiable_on {0..1} - s" by force qed have inj: "inj ?\" apply (intro comp_inj_on inj_on_cases inj_on_arccos) apply (auto simp add: inj_on_def x_coord_def) using pi_ge_zero apply auto[1] apply (smt arccos) by (smt arccos_lbound) then have fin: "\x. \0 \ x; x \ 1\ \ finite (?\ -` {x})" by (simp add: finite_vimageI) have "?\ ` {0..1} = {0..1}" proof show "?\ ` {0..1} \ {0..1}" by (auto simp add: x_coord_def divide_simps arccos_lbound arccos_bounded) have "arccos (1 - 2 * ((1 - cos (x * pi))/2)) = x * pi" if "0 \ x" "x \ 1" for x using that by (simp add: field_simps arccos_cos) then show "{0..1} \ ?\ ` {0..1}" apply (auto simp: x_coord_def o_def image_def) by (rule_tac x="(1 - cos (x * pi))/2" in bexI) auto qed then have bij_phi: "bij_betw ?\ {0..1} {0..1}" unfolding bij_betw_def using inj inj_on_subset by blast have phi01: "?\ -` {0..1} \ {0..1}" by (auto simp add: subset_iff x_coord_def divide_simps) { fix x::real assume x: "0 \ x" "x \ 1" then have i: "- 1 \ (2 * x - 1)" "(2 * x - 1) \ 1" by auto have ii: "cos (pi / 2 + arccos (1 - x * 2)) = -sin (arccos (1 - x * 2))" using minus_sin_cos_eq[symmetric, where ?x = "arccos (1 - x * 2)"] by (auto simp add: add.commute) have "2 * sqrt (x - x * x) = sqrt (4*x - 4*x * x)" by (metis mult.assoc real_sqrt_four real_sqrt_mult right_diff_distrib) also have "... = sqrt (1 - (2 * x - 1) * (2 * x - 1))" by (simp add: algebra_simps) finally have iii: "2 * sqrt (x - x * x) = cos (arcsin (2 * x - 1)) \ 2 * sqrt (x - x * x) = sin (arccos (1 - 2 * x))" using arccos_minus[OF i] unfolding minus_diff_eq sin_pi_minus by (simp add: cos_arcsin i power2_eq_square sin_arccos) then have 1: "?LHS1 x = ?RHS1 x" using d_gt_0 i x apply (simp add: x_coord_def field_simps) apply (auto simp add: diff_divide_distrib add_divide_distrib right_diff_distrib mult.commute ii) using cos_sin_eq[where ?x = "- arccos (1 - x * 2)"] by (simp add: cos_sin_eq[where ?x = "- arccos (1 - x * 2)"] right_diff_distrib) have 2: "?LHS2 x = ?RHS2 x" using x d_gt_0 iii by (auto simp add: x_coord_def diff_divide_distrib algebra_simps) have a: "cos (pi / 2 - arccos (x * 2 - 1)) = cos (pi / 2 - arccos (1 - x * 2))" by (smt arccos_minus arccos_cos2 arccos_lbound cos_arccos cos_ge_minus_one cos_le_one i(1) i(2) pi_def pi_half) have b: "cos (arccos (1 - x * 2) + pi * 3 / 2) = cos ((pi / 2 - arccos (x * 2 - 1)) + 2 * pi)" using arccos_minus[OF i] by(auto simp add: mult.commute add.commute) then have c: "... = cos (pi / 2 - arccos (x * 2 - 1))" using cos_periodic by blast have "cos (- pi - arccos (1 - x * 2)) = cos (- (pi + arccos (1 - x * 2)))" by (auto simp add: minus_add[where b = "pi" and a = "arccos (1 - x * 2)", symmetric]) moreover have "... = cos ((pi + arccos (1 - x * 2)))" using cos_minus by blast moreover have "... = cos (2*pi - arccos (x * 2 - 1))" using arccos_minus[OF i] by (auto simp add: mult.commute add.commute) moreover have "... = cos (arccos (x * 2 - 1))" using cos_2pi_minus by auto ultimately have d: "cos (- pi - arccos (1 - x * 2)) = (x * 2 - 1)" using cos_arccos[OF i] mult.commute by metis have cosm: "\x. cos (x - pi*2) = cos x" by (metis cos_periodic eq_diff_eq' mult.commute) have 34: "?LHS3 x = (?RHS3 \ ?\) x" "?LHS4 x = (?RHS4 \ ?\) x" using d_gt_0 x a b c iii cos_periodic [of "pi / 2 - arccos (x * 2 - 1)"] apply (auto simp add: x_coord_def algebra_simps diff_divide_distrib power2_eq_square) apply (auto simp add: sin_cos_eq cosm) using d apply (auto simp add: right_diff_distrib) by (smt cos_minus) note 1 2 34 } note * = this show ?P1 ?P2 ?P3 ?P4 apply (auto simp add: subcube_def circle_bot_edge_def circle_top_edge_def circle_polar_def reversepath_def subpath_def joinpaths_def circle_y_def rot_circle_left_edge_def rot_circle_right_edge_def) unfolding reparam_def by (rule ballI exI conjI impI phi_diff bij_phi phi01 fin * | force simp add: x_coord_def)+ qed definition circle_cube_boundary_to_polarcircle where "circle_cube_boundary_to_polarcircle \ \ if (\ = (circle_top_edge)) then {subcube 0 (1/4) (1, circle_polar), subcube (1/4) (1/2) (1, circle_polar)} else if (\ = (circle_bot_edge)) then {subcube (1/2) (3/4) (1, circle_polar), subcube (3/4) 1 (1, circle_polar)} else {}" definition rot_circle_cube_boundary_to_polarcircle where "rot_circle_cube_boundary_to_polarcircle \ \ if (\ = (rot_circle_left_edge )) then {subcube (1/4) (1/2) (1, circle_polar), subcube (1/2) (3/4) (1, circle_polar)} else if (\ = (rot_circle_right_edge)) then {subcube (3/4) 1 (1, circle_polar), subcube 0 (1/4) (1, circle_polar)} else {}" lemma circle_arcs_neq: assumes "0 \ k" "k \ 1" "0 \ n" "n \ 1" "n < k" "k + n < 1" shows "subcube k m (1, circle_polar) \ subcube n q (1, circle_polar)" proof (simp add: subcube_def subpath_def circle_polar_def) have "cos (2 * pi * k) \ cos(2 * pi * n)" unfolding cos_eq proof safe show False if "2 * pi * k = 2 * pi * n + 2 * m * pi" "m \ \" for m proof - have "2 * pi * (k - n ) = 2 * m * pi" using distrib_left that by (simp add: left_diff_distrib mult.commute) then have a: "m = (k - n)" by auto have "\k - n\ = 0 " using assms by (simp add: floor_eq_iff) then have "k - n \ \" using assms by (auto simp only: frac_eq_0_iff[symmetric] frac_def) then show False using that a by auto qed show False if "2 * pi * k = - (2 * pi * n) + 2 * m * pi" "m \ \" for m proof - have "2 * pi * (k + n ) = 2 * m * pi" using that by (auto simp add: distrib_left) then have a: "m = (k + n)" by auto have "\k + n\ = 0 " using assms by (simp add: floor_eq_iff) then have "k + n \ \" using Ints_def assms by force then show False using that a by auto qed qed then have "(\x. (d * cos (2 * pi * ((m - k) * x + k))/2, d * sin (2 * pi * ((m - k) * x + k))/2)) 0 \ (\x. (d * cos (2 * pi * ((q - n) * x + n))/2, d * sin (2 * pi * ((q - n) * x + n))/2)) 0" using d_gt_0 by auto then show "(\x. (d * cos (2 * pi * ((m - k) * x + k))/2, d * sin (2 * pi * ((m - k) * x + k))/2)) \ (\x. (d * cos (2 * pi * ((q - n) * x + n))/2, d * sin (2 * pi * ((q - n) * x + n))/2))" by metis qed lemma circle_arcs_neq_2: assumes "0 \ k" "k \ 1" "0 \ n" "n \ 1" "n < k" "0 < n" and kn12: "1/2 < k + n" and "k + n < 3/2" shows "subcube k m (1, circle_polar) \ subcube n q (1, circle_polar)" proof (simp add: subcube_def subpath_def circle_polar_def) have "sin (2 * pi * k) \ sin(2 * pi * n)" unfolding sin_eq proof safe show False if "m \ \" "2 * pi * k = 2 * pi * n + 2 * m * pi" for m proof - have "2 * pi * (k - n) = 2 * m * pi" using that by (simp add: left_diff_distrib mult.commute) then have a: "m = (k - n)" by auto have "\k - n\ = 0 " using assms by (simp add: floor_eq_iff) then have "k - n \ \" using assms by (auto simp only: frac_eq_0_iff[symmetric] frac_def ) then show False using that a by auto qed show False if "2 * pi * k = - (2 * pi * n) + (2 * m + 1) * pi" "m \ \" for m proof - have i: "\pi. 0 < pi \ 2 * pi * (k + n ) = 2 * m * pi + pi \ m = (k + n) - 1/2" by (sos "(((((A<0 * A<1) * R<1) + ([1/2] * A=0))) & ((((A<0 * A<1) * R<1) + ([~1/2] * A=0))))") have "2 * pi * (k + n) = 2 * m * pi + pi" using that by (simp add: algebra_simps) then have a: "m = (k + n) - 1/2" using i[OF pi_gt_zero] by fastforce have "\k + n - 1/2\ = 0 " using assms by (auto simp add: floor_eq_iff) then have "k + n - 1/2 \ \" by (metis Ints_cases add.commute add.left_neutral add_diff_cancel_left' add_diff_eq kn12 floor_of_int of_int_0 order_less_irrefl) then show False using that a by auto qed qed then have "(\x. (d * cos (2 * pi * ((m - k) * x + k))/2, d * sin (2 * pi * ((m - k) * x + k))/2)) 0 \ (\x. (d * cos (2 * pi * ((q - n) * x + n))/2, d * sin (2 * pi * ((q - n) * x + n))/2)) 0" using d_gt_0 by auto then show "(\x. (d * cos (2 * pi * ((m - k) * x + k))/2, d * sin (2 * pi * ((m - k) * x + k))/2)) \ (\x. (d * cos (2 * pi * ((q - n) * x + n))/2, d * sin (2 * pi * ((q - n) * x + n))/2))" by metis qed lemma circle_cube_is_only_horizontal_div_of_rot: shows "only_horizontal_division (boundary (circle_cube)) {rot_circle_cube}" unfolding only_horizontal_division_def proof (rule exI [of _ "{}"], simp, intro conjI ballI) show "finite (boundary circle_cube)" using circle.circle_cube_boundary_explicit circle_axioms by auto show "boundary_chain (boundary circle_cube)" by (simp add: two_cube_boundary_is_boundary) show "\x. x \ boundary circle_cube \ case x of (k, x) \ valid_path x" using circle_cube_boundary_valid by blast let ?\ = "(boundary (circle_cube))" let ?pi = "{circle_left_edge, circle_right_edge}" let ?pj = "{rot_circle_top_edge, rot_circle_bot_edge}" let ?f = "circle_cube_boundary_to_polarcircle" let ?one_chaini = "boundary (circle_cube) - ?pi" have c: "common_reparam_exists ?\ (two_chain_vertical_boundary {rot_circle_cube})" unfolding common_reparam_exists_def proof (intro exI conjI) let ?subdiv = "{(subcube 0 (1/4) (1, circle_polar)), (subcube (1/4) (1/2) (1, circle_polar)), (subcube (1/2) (3/4) (1, circle_polar)), (subcube (3/4) 1 (1, circle_polar))}" show "(\(k, \)\?subdiv. \ C1_differentiable_on {0..1})" using subpath_smooth[OF circle_polar_smooth] by (auto simp add: subcube_def) have 1: "finite ?subdiv" by auto show "boundary_chain ?subdiv" by (simp add: boundary_chain_def subcube_def) show "chain_reparam_chain' (boundary (circle_cube) - ?pi) ?subdiv" unfolding chain_reparam_chain'_def proof (intro exI conjI impI) show "\ (?f ` ?one_chaini) = ?subdiv" apply (simp add: circle_cube_boundary_to_polarcircle_def circle_cube_boundary_explicit) using circle_top_bot_edges_neq' by metis let ?l = "[subcube 0 (1/4) (1, circle_polar), subcube (1/4) (1/2) (1, circle_polar)]" have "chain_reparam_weak_path (coeff_cube_to_path (circle_top_edge)) {subcube 0 (1/4) (1, circle_polar), subcube (1/4) (1/2) (1, circle_polar)}" unfolding chain_reparam_weak_path_def proof (intro exI conjI) show "valid_chain_list ?l" by (auto simp add: subcube_def circle_top_edge_def x_coord_def circle_y_def pathfinish_def pathstart_def reversepath_def subpath_def joinpaths_def) show "reparam (coeff_cube_to_path circle_top_edge) (rec_join ?l)" using circ_top_edge_reparam_polar_circ_split by auto show "distinct ?l" apply simp apply (subst neq_commute) apply (simp add: circle_arcs_neq) done qed auto moreover have "chain_reparam_weak_path (coeff_cube_to_path (circle_bot_edge)) {subcube (1/2) (3/4) (1, circle_polar), subcube (3/4) 1 (1, circle_polar)}" unfolding chain_reparam_weak_path_def proof let ?l = "[subcube (1/2) (3/4) (1, circle_polar), subcube (3/4) 1 (1, circle_polar)]" have a: "valid_chain_list ?l" by (auto simp add: subcube_def circle_top_edge_def x_coord_def circle_y_def pathfinish_def pathstart_def reversepath_def subpath_def joinpaths_def) have b: "reparam (rec_join [circle_bot_edge]) (rec_join ?l)" using circ_bot_edge_reparam_polar_circ_split by auto have c: "subcube (3/4) 1 (1, circle_polar) \ subcube (1/2) (3/4) (1, circle_polar)" apply(rule circle_arcs_neq_2) using d_gt_0(1) by auto show "set ?l = {subcube (1/2) (3/4) (1, circle_polar), subcube (3/4) 1 (1, circle_polar)} \ distinct ?l \ reparam (coeff_cube_to_path (circle_bot_edge)) (rec_join ?l) \ valid_chain_list ?l \ ?l \ []" using a b c by auto qed ultimately show "(\cube\?one_chaini. chain_reparam_weak_path (rec_join [cube]) (?f cube))" by (auto simp add: circle_cube_boundary_to_polarcircle_def UNION_eq circle_cube_boundary_explicit) show "(\p\?one_chaini. \p'\?one_chaini. p \ p' \ ?f p \ ?f p' = {})" using circle_arcs_neq circle_arcs_neq_2 apply (auto simp add: circle_cube_boundary_to_polarcircle_def UNION_eq circle_cube_boundary_explicit neq_commute d_gt_0) using circle_top_bot_edges_neq' d_gt_0 apply auto[1] apply (smt atLeastAtMost_iff divide_less_eq_1_pos zero_less_divide_1_iff) apply (smt atLeastAtMost_iff divide_less_eq_1_pos zero_less_divide_iff) apply (smt atLeastAtMost_iff divide_cancel_left divide_less_eq_1_pos field_sum_of_halves zero_less_divide_1_iff) done show "(\x\?one_chaini. finite (?f x))" by (auto simp add: circle_cube_boundary_to_polarcircle_def circle_cube_boundary_explicit) qed show "(\(k, \)\?pi. point_path \)" using d_gt_0 by (auto simp add: point_path_def circle_left_edge_def circle_right_edge_def) let ?f = "rot_circle_cube_boundary_to_polarcircle" let ?one_chain2.0 = "two_chain_vertical_boundary {rot_circle_cube} - ?pj" show "chain_reparam_chain' (two_chain_vertical_boundary {rot_circle_cube} - ?pj) ?subdiv" unfolding chain_reparam_chain'_def proof (intro exI conjI) have rw: "?one_chain2.0 = {rot_circle_left_edge, rot_circle_right_edge}" by(auto simp add: rot_circle_cube_vertical_boundary_explicit two_chain_vertical_boundary_def) show "\ (?f ` ?one_chain2.0) = ?subdiv" using rot_circle_right_top_edges_neq' by (auto simp add: rot_circle_cube_boundary_to_polarcircle_def rw) show "(\cube\?one_chain2.0. chain_reparam_weak_path (rec_join [cube]) (?f cube))" proof (clarsimp simp add: rot_circle_cube_boundary_to_polarcircle_def rw, intro conjI) let ?l = "[subcube (1/4) (1/2) (1, circle_polar), subcube (1/2) (3/4) (1, circle_polar)]" show "chain_reparam_weak_path (coeff_cube_to_path (rot_circle_left_edge)) {subcube (1/4) (1/2) (1, circle_polar), subcube (1/2) (3/4) (1, circle_polar)}" unfolding chain_reparam_weak_path_def proof (intro exI conjI) show "valid_chain_list ?l" by (auto simp add: subcube_def pathfinish_def pathstart_def reversepath_def subpath_def joinpaths_def) show "reparam (coeff_cube_to_path rot_circle_left_edge) (rec_join ?l)" using rot_circ_left_edge_reparam_polar_circ_split by auto show "distinct ?l" apply simp apply (subst neq_commute) apply (simp add: circle_arcs_neq) done qed auto show "chain_reparam_weak_path (coeff_cube_to_path (rot_circle_right_edge)) {subcube (3/4) 1 (1, circle_polar), subcube 0 (1/4) (1, circle_polar)}" unfolding chain_reparam_weak_path_def proof (intro exI conjI) let ?l = "[subcube (3/4) 1 (1, circle_polar), subcube 0 (1/4) (1, circle_polar)]" show "valid_chain_list ?l" by (auto simp add: circle_polar_def subcube_def pathfinish_def pathstart_def reversepath_def subpath_def joinpaths_def) show "reparam (coeff_cube_to_path rot_circle_right_edge) (rec_join ?l)" using rot_circ_right_edge_reparam_polar_circ_split by auto show "distinct ?l" by (simp add: circle_arcs_neq) qed auto qed show "(\p\?one_chain2.0. \p'\?one_chain2.0. p \ p' \ ?f p \ ?f p' = {})" using circle_arcs_neq circle_arcs_neq_2 apply (auto simp add: rot_circle_cube_boundary_to_polarcircle_def neq_commute) apply (metis add.right_neutral divide_less_eq_1_pos dual_order.order_iff_strict num.distinct(1) one_less_numeral_iff prod.sel(1) prod.sel(2) semiring_norm(68) subcube_def zero_less_divide_1_iff zero_less_numeral) apply (smt field_sum_of_halves) done show "(\x\?one_chain2.0. finite (?f x))" by (auto simp add: rot_circle_cube_boundary_to_polarcircle_def UNION_eq rw) qed show "(\(k, \)\?pj. point_path \)" using d_gt_0 by (auto simp add: point_path_def rot_circle_top_edge_def rot_circle_bot_edge_def) qed then show "common_sudiv_exists (two_chain_vertical_boundary {rot_circle_cube}) (boundary circle_cube) \ common_reparam_exists (boundary circle_cube) (two_chain_vertical_boundary {rot_circle_cube})" by blast qed lemma GreenThm_cirlce: assumes "\twoC\{circle_cube}. analytically_valid (cubeImage twoC) (\x. F x \ i) j" "\twoC\{rot_circle_cube}. analytically_valid (cubeImage twoC) (\x. F x \ j) i" shows "integral (cubeImage (circle_cube)) (\x. partial_vector_derivative (\x. F x \ j) i x - partial_vector_derivative (\x. F x \ i) j x) = one_chain_line_integral F {i, j} (boundary (circle_cube))" proof(rule green_typeI_typeII_chain.GreenThm_typeI_typeII_divisible_region_finite_holes[of "(cubeImage (circle_cube))" i j F "{circle_cube}" "{rot_circle_cube}", OF _ _ _ circle_cube_is_only_horizontal_div_of_rot _], auto) show "\ a b. (a, b) \ boundary circle_cube \ valid_path b" using circle_cube_boundary_valid by auto show "green_typeI_typeII_chain (cubeImage circle_cube) i j F {circle_cube} {rot_circle_cube}" using assms proof(auto simp add: green_typeI_typeII_chain_def green_typeI_chain_def green_typeII_chain_def green_typeI_chain_axioms_def green_typeII_chain_axioms_def intro!: circle_cube_is_type_I rot_circle_cube_is_type_II d_gt_0 R2_axioms) show "gen_division (cubeImage circle_cube) {cubeImage circle_cube}" by (simp add: gen_division_def) show "gen_division (cubeImage (circle_cube)) ({cubeImage rot_circle_cube})" using rot_circle_div_circle d_gt_0 by auto show "valid_two_chain {rot_circle_cube}" "valid_two_chain {circle_cube}" apply (auto simp add: valid_two_chain_def) using rot_circle_cube_valid_two_cube circle_cube_valid_two_cube assms(1) by auto qed show "only_vertical_division (boundary (circle_cube)) {circle_cube}" using twoChainVertDiv_of_itself[of "{circle_cube}"] apply(simp add: two_chain_boundary_def) using circle_cube_boundary_valid by auto qed end end diff --git a/thys/LLL_Factorization/LLL_Factorization.thy b/thys/LLL_Factorization/LLL_Factorization.thy --- a/thys/LLL_Factorization/LLL_Factorization.thy +++ b/thys/LLL_Factorization/LLL_Factorization.thy @@ -1,1328 +1,1322 @@ (* Authors: Jose Divasón Sebastiaan Joosten René Thiemann Akihisa Yamada License: BSD *) section \Correctness of the LLL factorization algorithm\ text \This theory connects short vectors of lattices and factors of polynomials. From this connection, we derive soundness of the lattice based factorization algorithm. \ theory LLL_Factorization imports LLL_Factorization_Impl Berlekamp_Zassenhaus.Factorize_Int_Poly begin subsection \Basic facts about the auxiliary functions\ hide_const (open) module.smult lemma nth_factorization_lattice: fixes u and d defines "n \ degree u" assumes "i < n + d" shows "factorization_lattice u d m ! i = vec_of_poly_n (if i < d then u * monom 1 (d - Suc i) else monom m (n+d-Suc i)) (n+d)" using assms by (unfold factorization_lattice_def, auto simp: nth_append smult_monom Let_def not_less) lemma length_factorization_lattice[simp]: shows "length (factorization_lattice u d m) = degree u + d" by (auto simp: factorization_lattice_def Let_def) lemma dim_factorization_lattice: assumes "x < degree u + d" shows "dim_vec (factorization_lattice u d m ! x) = degree u + d" unfolding factorization_lattice_def using assms nth_append by (simp add: nth_append Let_def) lemma dim_factorization_lattice_element: assumes "x \ set (factorization_lattice u d m)" shows "dim_vec x = degree u + d" using assms by (auto simp: factorization_lattice_def Let_def) lemma set_factorization_lattice_in_carrier[simp]: "set (factorization_lattice u d m) \ carrier_vec (degree u + d)" using dim_factorization_lattice by (auto simp: factorization_lattice_def Let_def) lemma choose_u_Cons: "choose_u (x#xs) = (if xs = [] then x else min_degree_poly x (choose_u xs))" by (cases xs, auto) lemma choose_u_member: "xs \ [] \ choose_u xs \ set xs" by (induct xs, auto simp: choose_u_Cons) declare choose_u.simps[simp del] subsection \Facts about Sylvester matrices and norms\ lemma (in LLL) lattice_is_span [simp]: "lattice_of xs = span_list xs" by (unfold lattice_of_def span_list_def lincomb_list_def image_def, auto) lemma sq_norm_row_sylvester_mat1: fixes f g :: "'a :: conjugatable_ring poly" assumes i: "i < degree g" shows "\(row (sylvester_mat f g) i)\\<^sup>2 = \f\\<^sup>2" proof (cases "f = 0") case True thus ?thesis by (auto simp add: sylvester_mat_def row_def sq_norm_vec_def o_def interv_sum_list_conv_sum_set_nat i intro!: sum_list_zero) next case False note f = False let ?f = "\j. if i \ j \ j - i \ degree f then coeff f (degree f + i - j) else 0" let ?h = "\j. j + i" let ?row = "vec (degree f + degree g) ?f " let ?g = "\j. degree f - j" have image_g: "?g ` {0..row (sylvester_mat f g) i\\<^sup>2 = \?row\\<^sup>2" by (rule arg_cong[of _ _ "sq_norm_vec"], insert i, auto simp add: row_def sylvester_mat_def sylvester_mat_sub_def) also have "... = sum_list (map (sq_norm \ ?f) [0.. ?f) {0.. ?f) {i..< Suc (degree f + i)}" by (rule sum.mono_neutral_right, insert i, auto) also have "... = sum ((sq_norm \ ?f) \ ?h) {0..j. sq_norm (coeff f (degree f - j))) {0..j. sq_norm (coeff f j)) \ ?g) {0..j. sq_norm (coeff f j)) (?g ` {0.. coeff f) {0.. i" and i2: "i < degree f + degree g" shows "\row (sylvester_mat f g) i\\<^sup>2 = \g\\<^sup>2" proof - let ?f = "\j. if i - degree g \ j \ j \ i then coeff g (i - j) else 0" let ?row = "vec (degree f + degree g) ?f" let ?h = "\j. j + i - degree g" let ?g = "\j. degree g - j" have image_g: "?g ` {0.. degree g" if x: "x < Suc i" for x using x by auto have bij_h: "bij_betw ?h {0..row (sylvester_mat f g) i\\<^sup>2 = \?row\\<^sup>2" by (rule arg_cong[of _ _ "sq_norm_vec"], insert i1 i2, auto simp add: row_def sylvester_mat_def sylvester_mat_sub_def) also have "... = sum_list (map (sq_norm \ ?f) [0.. ?f) {0.. ?f) {i - degree g..< Suc i}" by (rule sum.mono_neutral_right, insert i2, auto) also have "... = sum ((sq_norm \ ?f) \ ?h) {0..j. sq_norm (coeff g (degree g - j))) {0..j. sq_norm (coeff g j)) \ ?g) {0..j. sq_norm (coeff g j)) (?g ` {0.. coeff g) {0.. carrier_mat n n" shows "\det A\ \ sqrt (of_int (prod_list (map sq_norm (rows A))))" proof - let ?A = "map_mat real_of_int A" have "\det A\ = \det ?A\" unfolding of_int_hom.hom_det by simp also have "\ \ sqrt (prod_list (map sq_norm (rows ?A)))" by (rule Hadamard's_inequality[of ?A n], insert A, auto) also have "\ = sqrt (of_int (prod_list (map sq_norm (rows A))))" unfolding of_int_hom.hom_prod_list map_map by (rule arg_cong[of _ _ "\ x. sqrt (prod_list x)"], rule nth_equalityI, force, auto simp: sq_norm_of_int[symmetric] row_def intro!: arg_cong[of _ _ sq_norm_vec]) finally show ?thesis . qed lemma resultant_le_prod_sq_norm: fixes f g::"int poly" defines "n \ degree f" and "k \ degree g" shows "\resultant f g\ \ sqrt (of_int (\f\\<^sup>2^k * \g\\<^sup>2^n))" proof - let ?S = "sylvester_mat f g" let ?f = "sq_norm \ row ?S" have map_rw1: "map ?f [0..f\\<^sup>2" proof (rule nth_equalityI) let ?M = "map (sq_norm \ row (sylvester_mat f g)) [0..f\\<^sup>2)" using k_def by auto show "?M ! i = replicate k \f\\<^sup>2 ! i" if i: "i < length ?M" for i proof - have ik: "if\\<^sup>2 ! i = \f\\<^sup>2" by (rule nth_replicate[OF ik]) also have "... = (sq_norm \ row (sylvester_mat f g)) (0 + i)" using sq_norm_row_sylvester_mat1 ik k_def by force also have "... = ?M ! i" by (rule nth_map_upt[symmetric], simp add: i_deg_g) finally show "?M ! i = replicate k \f\\<^sup>2 ! i" .. qed qed have map_rw2: "map ?f [degree g..g\\<^sup>2" proof (rule nth_equalityI) let ?M = "map (sq_norm \ row (sylvester_mat f g)) [degree g..g\\<^sup>2)" by (simp add: n_def) show "?M ! i = replicate n \g\\<^sup>2 ! i" if "ig\\<^sup>2 ! i = \g\\<^sup>2" by (rule nth_replicate[OF i_n]) also have "... = (sq_norm \ row (sylvester_mat f g)) (degree g + i)" using i_n n_def by (simp add: sq_norm_row_sylvester_mat2) also have "... = ?M ! i" by (simp add: i_deg_f) finally show "?M ! i = replicate n \g\\<^sup>2 ! i" .. qed qed have p1: "prod_list (map ?f [0..f\\<^sup>2^k" unfolding map_rw1 by (rule prod_list_replicate) have p2: "prod_list (map ?f [degree g..g\\<^sup>2^n" unfolding map_rw2 by (rule prod_list_replicate) have list_rw: "[0..resultant f g\ = \det ?S\" unfolding resultant_def .. also have "... \ sqrt (of_int (prod_list (map sq_norm (rows ?S))))" by (rule Hadamard's_inequality_int, auto) also have "map sq_norm (rows ?S) = map ?f [0..Proof of the key lemma 16.20\ (* Lemma 16.20 *) lemma common_factor_via_short: fixes f g u :: "int poly" defines "n \ degree f" and "k \ degree g" assumes n0: "n > 0" and k0: "k > 0" and monic: "monic u" and deg_u: "degree u > 0" and uf: "poly_mod.dvdm m u f" and ug: "poly_mod.dvdm m u g" and short: "\f\\<^sup>2^k * \g\\<^sup>2^n < m\<^sup>2" and m: "m \ 0" shows "degree (gcd f g) > 0" proof - interpret poly_mod m . have f_not0: "f \ 0" and g_not0: "g \ 0" using n0 k0 k_def n_def by auto have deg_f: "degree f > 0" using n0 n_def by simp have deg_g: "degree g > 0" using k0 k_def by simp obtain s t where deg_s: "degree s < degree g" and deg_t: "degree t < degree f" and res_eq: "[:resultant f g:] = s * f + t * g" and s_not0: "s \ 0" and t_not0: "t \ 0" using resultant_as_nonzero_poly[OF deg_f deg_g] by auto have res_eq_modulo: "[:resultant f g:] =m s * f + t * g" using res_eq by simp have u_dvdm_res: "u dvdm [:resultant f g:]" proof (unfold res_eq, rule dvdm_add) show "u dvdm s * f" using dvdm_factor[OF uf, of s] unfolding mult.commute[of f s] by auto show "u dvdm t * g" using dvdm_factor[OF ug, of t] unfolding mult.commute[of g t] by auto qed have res_0_mod: "resultant f g mod m = 0" by (rule monic_dvdm_constant[OF u_dvdm_res monic deg_u]) have res0: "resultant f g = 0" proof (rule mod_0_abs_less_imp_0) show "[resultant f g = 0] (mod m)" using res_0_mod unfolding cong_def by auto - have "\resultant f g\ \ sqrt ((sq_norm_poly f)^k * (sq_norm_poly g)^n)" + have "\resultant f g\ \ sqrt (real_of_int (\f\\<^sup>2 ^ k * \g\\<^sup>2 ^ n))" unfolding k_def n_def by (rule resultant_le_prod_sq_norm) also have "... < m" - proof (rule real_less_lsqrt) - show "0 \ real_of_int (\f\\<^sup>2 ^ k * \g\\<^sup>2 ^ n)" - by (simp add: sq_norm_poly_ge_0) - show "0 \ real_of_int m" using m by simp - show "real_of_int (\f\\<^sup>2 ^ k * \g\\<^sup>2 ^ n) < (real_of_int m)\<^sup>2" - by (metis of_int_less_iff of_int_power short) - qed + by (meson m of_int_0_le_iff of_int_power_less_of_int_cancel_iff real_less_lsqrt short) finally show "\resultant f g\ < m" using of_int_less_iff by blast qed have "\ coprime f g" by (rule resultant_zero_imp_common_factor, auto simp add: deg_f res0) thus ?thesis using res0 resultant_0_gcd by auto qed subsection \Properties of the computed lattice and its connection with Sylvester matrices\ lemma factorization_lattice_as_sylvester: fixes p :: "'a :: semidom poly" assumes dj: "d \ j" and d: "degree p = d" shows "mat_of_rows j (factorization_lattice p (j-d) m) = sylvester_mat_sub d (j-d) p [:m:]" proof (cases "p=0") case True have deg_p: "d = 0" using True d by simp show ?thesis by (auto simp add: factorization_lattice_def True deg_p mat_of_rows_def d) next case p0: False note 1 = degree_mult_eq[OF p0, of "monom _ _", unfolded monom_eq_0_iff, OF one_neq_zero] from dj show ?thesis apply (cases "m = 0") apply (auto simp: mat_eq_iff d[symmetric] 1 coeff_mult_monom sylvester_mat_sub_index mat_of_rows_index nth_factorization_lattice vec_index_of_poly_n degree_monom_eq coeff_const) done qed context inj_comm_semiring_hom begin lemma map_poly_hom_mult_monom [hom_distribs]: "map_poly hom (p * monom a n) = map_poly hom p * monom (hom a) n" by (auto intro!: poly_eqI simp:coeff_mult_monom hom_mult) lemma hom_vec_of_poly_n [hom_distribs]: "map_vec hom (vec_of_poly_n p n) = vec_of_poly_n (map_poly hom p) n" by (auto simp: vec_index_of_poly_n) lemma hom_factorization_lattice [hom_distribs]: shows "map (map_vec hom) (factorization_lattice u k m) = factorization_lattice (map_poly hom u) k (hom m)" by (auto intro!:arg_cong[of _ _ "\p. vec_of_poly_n p _"] simp: list_eq_iff_nth_eq nth_factorization_lattice hom_vec_of_poly_n map_poly_hom_mult_monom) end subsection \Proving that @{const factorization_lattice} returns a basis of the lattice\ context LLL begin sublocale idom_vec n "TYPE(int)". (*In this context, "n" is fixed by the locale and corresponds to "j" in the book*) lemma upper_triangular_factorization_lattice: fixes u :: "'a :: semidom poly" and d :: nat assumes d: "d \ n" and du: "d = degree u" shows "upper_triangular (mat_of_rows n (factorization_lattice u (n-d) k))" (is "upper_triangular ?M") proof (intro upper_triangularI, unfold mat_of_rows_carrier length_factorization_lattice) fix i j assume ji: "j < i" and i: "i < degree u + (n - d)" with d du have jn: "j < n" by auto show "?M $$ (i,j) = 0" proof (cases "u=0") case True with ji i show ?thesis by (auto simp: factorization_lattice_def mat_of_rows_def) next case False then show ?thesis using d ji i apply (simp add: du mat_of_rows_index nth_factorization_lattice) apply (auto simp: vec_index_of_poly_n[OF jn] degree_mult_eq degree_monom_eq) done qed qed lemma factorization_lattice_diag_nonzero: fixes u :: "'a :: semidom poly" and d assumes d: "d=degree u" and dn: "d\n" and u: "u\0" and m0: "k\0" and i: "i 0" proof- have 1: "monom (1::'a) (n - Suc (degree u + i)) \ 0" using m0 by auto have 2: "i < degree u + (n - d)" using i d by auto let ?p = "u * monom 1 (n - Suc (degree u + i))" have 3: "i < n - degree u \ degree (?p) = n - Suc i" using assms by (auto simp: degree_mult_eq[OF _ 1] degree_monom_eq) show ?thesis apply (unfold nth_factorization_lattice[OF 2] vec_index_of_poly_n[OF 2]) using assms leading_coeff_0_iff[of ?p] apply (cases "i < n - degree u", auto simp: d 3 degree_monom_eq) done qed corollary factorization_lattice_diag_nonzero_RAT: fixes d assumes "d=degree u" and "d\n" and "u\0" and "k\0" and "i 0" using factorization_lattice_diag_nonzero[OF assms] assms by (auto simp: nth_factorization_lattice) sublocale gs: vec_space "TYPE(rat)" n. lemma lin_indpt_list_factorization_lattice: fixes d assumes d: "d = degree u" and dn: "d \ n" and u: "u \ 0" and k: "k \ 0" shows "gs.lin_indpt_list (RAT (factorization_lattice u (n-d) k))" (is "gs.lin_indpt_list (RAT ?vs)") proof- have 1: "rows (mat_of_rows n (map (map_vec rat_of_int) ?vs)) = map (map_vec rat_of_int) ?vs" using dn d by (subst rows_mat_of_rows, auto dest!: subsetD[OF set_factorization_lattice_in_carrier]) note 2 = factorization_lattice_diag_nonzero_RAT[OF d dn u k] show ?thesis apply (intro gs.upper_triangular_imp_lin_indpt_list[of "mat_of_rows n (RAT ?vs)", unfolded 1]) using assms 2 by (auto simp: diag_mat_def mat_of_rows_index hom_distribs intro!:upper_triangular_factorization_lattice) qed end subsection \Being in the lattice is being a multiple modulo\ lemma (in semiring_hom) hom_poly_of_vec: "map_poly hom (poly_of_vec v) = poly_of_vec (map_vec hom v)" by (auto simp add: coeff_poly_of_vec poly_eq_iff) abbreviation "of_int_vec \ map_vec of_int" context LLL begin lemma lincomb_to_dvd_modulo: fixes u d defines "d \ degree u" assumes d: "d \ n" and lincomb: "lincomb_list c (factorization_lattice u (n-d) k) = g" (is "?l = ?r") shows "poly_mod.dvdm k u (poly_of_vec g)" proof- let ?S = "sylvester_mat_sub d (n - d) u [:k:]" define q where "q \ poly_of_vec (vec_first (vec n c) (n - d))" define r where "r \ poly_of_vec (vec_last (vec n c) d)" have "?l = ?S\<^sup>T *\<^sub>v vec n c" apply (subst lincomb_list_as_mat_mult) using d d_def apply (force simp:factorization_lattice_def) apply (fold transpose_mat_of_rows) using d d_def by (simp add: factorization_lattice_as_sylvester) also have "poly_of_vec \ = q * u + smult k r" apply (subst sylvester_sub_poly) using d_def d q_def r_def by auto finally have "\ = poly_of_vec g" unfolding lincomb of_int_hom.hom_poly_of_vec by auto then have "poly_of_vec g = q * u + Polynomial.smult k r" by auto then have "poly_mod.Mp k (poly_of_vec g) = poly_mod.Mp k (q * u + Polynomial.smult k r)" by auto also have "... = poly_mod.Mp k (q * u + poly_mod.Mp k (Polynomial.smult k r))" using poly_mod.plus_Mp(2) by auto also have "... = poly_mod.Mp k (q * u)" using poly_mod.plus_Mp(2) unfolding poly_mod.Mp_smult_m_0 by simp also have "... = poly_mod.Mp k (u * q)" by (simp add: mult.commute) finally show ?thesis unfolding poly_mod.dvdm_def by auto qed (* There is a typo in the textbook (page 476). The book states q' = q'' * u + r'' and the correct fact is r' = q'' * u+r'' *) lemma dvd_modulo_to_lincomb: fixes u :: "int poly" and d defines "d \ degree u" assumes d: "d < n" and dvd: "poly_mod.dvdm k u (poly_of_vec g)" and k_not0: "k\0" and monic_u: "monic u" and dim_g: "dim_vec g = n" and deg_u: "degree u > 0" shows "\c. lincomb_list c (factorization_lattice u (n-d) k) = g" proof - interpret p: poly_mod k . have u_not0: "u \ 0" using monic_u by auto hence n[simp]: "0 < n" using d by auto obtain q' r' where g: "poly_of_vec g = q' * u + smult k r'" using p.dvdm_imp_div_mod[OF dvd] by auto obtain q'' r'' where r': "r' = q'' * u + r''" and deg_r'': "degree r'' degree (q' + smult k q'') < n - d" proof (cases "q = 0",auto, rule degree_div_mod_smult[OF _ _ _ g1]) show "degree (poly_of_vec g) < n" by (rule degree_poly_of_vec_less, auto simp add: dim_g) show "degree r'' < d" using deg_r'' unfolding d_def . assume "q\0" thus "q' + smult k q'' \ 0" unfolding q . show "k \ 0" by fact show "degree u = d" using d_def by auto qed have g2: "(vec_of_poly_n (q*u) n) + (vec_of_poly_n (smult k r) n) = g" proof - have "g = vec_of_poly_n (poly_of_vec g) n" by (rule vec_of_poly_n_poly_of_vec[symmetric], auto simp add: dim_g) also have "\ = vec_of_poly_n ((q' + smult k q'') * u + smult k r'') n" using g1 by auto also have "... = vec_of_poly_n (q * u + smult k r'') n" unfolding q by auto also have "... = vec_of_poly_n (q * u) n + vec_of_poly_n (smult k r'') n" by (rule vec_of_poly_n_add) finally show ?thesis unfolding r by simp qed let ?c = "\i. if i < n - d then coeff q (n - d - 1 - i) else coeff r (n - Suc i)" let ?c1 = "\i. ?c i \\<^sub>v factorization_lattice u (n-d) k ! i" show ?thesis proof (rule exI[of _ ?c]) let ?part1 = "map (\i. vec_of_poly_n (u * monom 1 i) n) [n-d>..0]" let ?part2 = "map (\i. vec_of_poly_n (monom k i) n) [d>..0]" have [simp]: "dim_vec (M.sumlist (map ?c1 [0.. carrier_vec n" if x: "x < n" for x using x dim_factorization_lattice_element nth_factorization_lattice[of x u "n-d"] d by (auto simp: d_def) have "[0..\<^bsub>V\<^esub>i\{0..\<^bsub>V\<^esub>i\set [0..i. poly_of_vec (?c1 i)) (set [0..i. poly_of_vec (?c1 i)) {0.. 0" using i by auto have deg_m: "degree ?m = n - Suc (d + i)" by (rule degree_monom_eq, auto) have "degree (u * ?m) = d + (n - Suc (d + i))" using degree_mult_eq[OF u_not0 monom_not0] d_def deg_m by auto also have "... < n" using i by auto finally show ?thesis . qed have lattice_rw: "factorization_lattice u (n-d) k ! i = vec_of_poly_n (u * monom 1 (n - Suc (d + i))) n" if i: "i< n - d" for i apply (subst nth_factorization_lattice) using i by (auto simp:d_def) have q_rw: "q = (\i = 0..x \ {0..x = 0..x \ insert ?m ({0..x \ {0..degree q" using degree_q q False d by auto then show ?thesis using coeff_eq_0 by auto qed finally show "coeff q j = (\i = 0..i. poly_of_vec (?c1 i)) {0..i = 0..\<^sub>v factorization_lattice u (n-d) k ! i))" by (rule sum.cong, auto) also have "... = (\i = 0..\<^sub>v (vec_of_poly_n (u * monom 1 (n - Suc (d + i))) n))))" by (rule sum.cong, auto simp add: lattice_rw) also have "... = (\i = 0..i = 0..i = 0..\<^bsub>V\<^esub>i\{n-d..\<^bsub>V\<^esub>i\set [n-d..i. poly_of_vec (?c1 i)) {n-d..i" and i2: "i i" and i2: "ii \ {n-d..i = n - d..i = n - d..i. poly_of_vec (?c1 i)) {n-d..i \ {n-d..\<^sub>v factorization_lattice u (n-d) k ! i))" by (rule sum.cong, auto) also have "... = (\i \ {n-d..\<^sub>v vec_of_poly_n (monom k (n - Suc i)) n)))" by (rule sum.cong, auto simp add: lattice_rw) also have "... = (\i \ {n-d..i \ {n-d..i \ {n-d..The factorization lattice precisely characterises the polynomials of a certain degree which divide $u$ modulo $M$.\ lemma factorization_lattice: fixes M assumes deg_u: "degree u \ 0" and M: "M \ 0" shows "degree u \ n \ n \ 0 \ f \ poly_of_vec ` lattice_of (factorization_lattice u (n - degree u) M) \ degree f < n \ poly_mod.dvdm M u f" "monic u \ degree u < n \ degree f < n \ poly_mod.dvdm M u f \ f \ poly_of_vec ` lattice_of (factorization_lattice u (n - degree u) M)" proof - from deg_u have deg_u: "degree u > 0" by auto let ?L = "factorization_lattice u (n - degree u) M" { assume deg: "degree f < n" and dvd: "poly_mod.dvdm M u f" and mon: "monic u" and deg_u_lt: "degree u < n" define fv where "fv = vec n (\ i. (coeff f (n - Suc i)))" have f: "f = poly_of_vec fv" unfolding fv_def poly_of_vec_def Let_def using deg by (auto intro!: poly_eqI coeff_eq_0 simp: coeff_sum) have dim_fv: "dim_vec fv = n" unfolding fv_def by simp from dvd_modulo_to_lincomb[OF deg_u_lt _ M mon _ deg_u(1), of fv, folded f, OF dvd dim_fv] obtain c where gv: "fv = lincomb_list c ?L" by auto have "fv \ lattice_of ?L" unfolding gv lattice_is_span by (auto simp: in_span_listI) thus "f \ poly_of_vec ` lattice_of ?L" unfolding f by auto } moreover { assume "f \ poly_of_vec ` lattice_of ?L" and deg_u: "degree u \ n" and n: "n \ 0" then obtain fv where f: "f = poly_of_vec fv" and fv: "fv \ lattice_of ?L" by auto from in_span_listE[OF fv[unfolded lattice_is_span]] obtain c where fv: "fv = lincomb_list c ?L" by auto from lincomb_to_dvd_modulo[OF _ fv[symmetric]] deg_u f have dvd: "poly_mod.dvdm M u f" by auto have "set ?L \ carrier_vec n" unfolding factorization_lattice_def using deg_u by auto hence "fv \ carrier_vec n" unfolding fv by (metis lincomb_list_carrier) hence "degree f < n" unfolding f using degree_poly_of_vec_less[of fv n] using n by auto with dvd show "degree f < n \ poly_mod.dvdm M u f" by auto } qed end subsection \Soundness of the LLL factorization algorithm\ lemma LLL_short_polynomial: assumes deg_u_0: "degree u \ 0" and deg_le: "degree u \ n" and pl1: "pl > 1" and monic: "monic u" shows "degree (LLL_short_polynomial pl n u) < n" and "LLL_short_polynomial pl n u \ 0" and "poly_mod.dvdm pl u (LLL_short_polynomial pl n u)" and "degree u < n \ f \ 0 \ poly_mod.dvdm pl u f \ degree f < n \ \LLL_short_polynomial pl n u\\<^sup>2 \ 2 ^ (n - 1) * \f\\<^sup>2" proof - interpret poly_mod_2 pl by (unfold_locales, insert pl1, auto) from pl1 have pl0: "pl \ 0" by auto let ?d = "degree u" let ?u = "Mp u" let ?iu = "inv_Mp ?u" from Mp_inv_Mp_id[of ?u] have "?iu =m ?u" . also have "\ =m u" by simp finally have iu_u: "?iu =m u" by simp have degu[simp]: "degree ?u = degree u" using monic by simp have mon: "monic ?u" using monic by (rule monic_Mp) have "degree ?iu = degree ?u" unfolding inv_Mp_def by (rule degree_map_poly, unfold mon, insert mon pl1, auto simp: inv_M_def) with degu have deg_iu: "degree ?iu = degree u" by simp have mon_iu: "monic ?iu" unfolding deg_iu unfolding inv_Mp_def Mp_def inv_M_def M_def by (insert pl1, auto simp: coeff_map_poly monic) let ?L = "factorization_lattice ?iu (n - ?d) pl" let ?sv = "short_vector_hybrid 2 ?L" from deg_u_0 deg_le have n: "n \ 0" by auto from deg_u_0 have u0: "u \ 0" by auto have id: "LLL_short_polynomial pl n u = poly_of_vec ?sv" unfolding LLL_short_polynomial_def by blast have id': "\?sv\\<^sup>2 = \LLL_short_polynomial pl n u\\<^sup>2" unfolding id by simp interpret vec_module "TYPE(int)" n. interpret L: LLL n n "?L" 2 . from deg_le deg_iu have deg_iu_le: "degree ?iu \ n" by simp have len: "length ?L = n" unfolding factorization_lattice_def using deg_le deg_iu by auto from deg_u_0 deg_iu have deg_iu0: "degree ?iu \ 0" by auto hence iu0: "?iu \ 0" by auto from L.lin_indpt_list_factorization_lattice[OF refl deg_iu_le iu0 pl0] have *: "4/3 \ (2 :: rat)" "L.gs.lin_indpt_list (L.RAT ?L)" by (auto simp: deg_iu) interpret L: LLL_with_assms n n ?L 2 by (unfold_locales, insert *, auto simp: deg_iu deg_le) note short = L.short_vector_hybrid[OF refl n, unfolded id' L.L_def] from short(2) have mem: "LLL_short_polynomial pl n u \ poly_of_vec ` lattice_of ?L" unfolding id by auto note fact = L.factorization_lattice(1)[OF deg_iu0 pl0 deg_iu_le n, unfolded deg_iu, OF mem] show "degree (LLL_short_polynomial pl n u) < n" using fact by auto from fact have "?iu dvdm (LLL_short_polynomial pl n u)" by auto then obtain h where "LLL_short_polynomial pl n u =m ?iu * h" unfolding dvdm_def by auto also have "?iu * h =m Mp ?iu * h" unfolding mult_Mp by simp also have "Mp ?iu * h =m u * h" unfolding iu_u unfolding mult_Mp by simp finally show "u dvdm (LLL_short_polynomial pl n u)" unfolding dvdm_def by auto from short have sv1: "?sv \ carrier_vec n" by auto from short have "?sv \ 0\<^sub>v j" for j by auto thus "LLL_short_polynomial pl n u \ 0" unfolding id by simp assume degu: "degree u < n" and dvd: "u dvdm f" and degf: "degree f < n" and f0: "f \ 0" from dvd obtain h where "f =m u * h" unfolding dvdm_def by auto also have "u * h =m Mp u * h" unfolding mult_Mp by simp also have "Mp u * h =m Mp ?iu * h" unfolding iu_u by simp also have "Mp ?iu * h =m ?iu * h" unfolding mult_Mp by simp finally have dvd: "?iu dvdm f" unfolding dvdm_def by auto from degu deg_iu have deg_iun: "degree ?iu < n" by auto from L.factorization_lattice(2)[OF deg_iu0 pl0 mon_iu deg_iun degf dvd] have "f \ poly_of_vec ` lattice_of ?L" using deg_iu by auto then obtain fv where f: "f = poly_of_vec fv" and fv: "fv \ lattice_of ?L" by auto have norm: "\fv\\<^sup>2 = \f\\<^sup>2" unfolding f by simp have fv0: "fv \ 0\<^sub>v n" using f0 unfolding f by auto with fv have fvL: "fv \ lattice_of ?L - {0\<^sub>v n}" by auto from short(3)[OF this, unfolded norm] have "rat_of_int \LLL_short_polynomial pl n u\\<^sup>2 \ rat_of_int (2 ^ (n - 1) * \f\\<^sup>2)" by simp thus "\LLL_short_polynomial pl n u\\<^sup>2 \ 2 ^ (n - 1) * \f\\<^sup>2" by linarith qed context LLL_implementation begin lemma LLL_reconstruction: assumes "LLL_reconstruction f us = fs" and "degree f \ 0" and "poly_mod.unique_factorization_m pl f (lead_coeff f, mset us)" and "f dvd F" and "\ ui. ui \ set us \ poly_mod.Mp pl ui = ui" and F0: "F \ 0" and cop: "coprime (lead_coeff F) p" and sf: "poly_mod.square_free_m p F" and pl1: "pl > 1" and plp: "pl = p^l" and p: "prime p" and large: "2^(5 * (degree F - 1) * (degree F - 1)) * \F\\<^sup>2^(2 * (degree F - 1)) < pl\<^sup>2" shows "f = prod_list fs \ (\ fi \ set fs. irreducible\<^sub>d fi)" proof - interpret p: poly_mod_prime p by (standard, rule p) interpret pl: poly_mod_2 "pl" by (standard, rule pl1) from pl1 plp have l0: "l \ 0" by (cases l, auto) show ?thesis using assms(1-5) proof (induct f us arbitrary: fs rule: LLL_reconstruction.induct) case (1 f us fs) define u where "u = choose_u us" define g where "g = LLL_short_polynomial pl (degree f) u" define k where "k = gcd f g" note res = 1(3) note degf = 1(4) note uf = 1(5) note fF = 1(6) note norm = 1(7) note to_fact = pl.unique_factorization_m_imp_factorization note fact = to_fact[OF uf] have mon_gs: "ui \ set us \ monic ui" for ui using norm fact unfolding pl.factorization_m_def by auto from p.coprime_lead_coeff_factor[OF p.prime] fF cop have cop: "coprime (lead_coeff f) p" unfolding dvd_def by blast have plf0: "pl.Mp f \ 0" using fact pl.factorization_m_lead_coeff pl.unique_factorization_m_zero uf by fastforce have "degree f = pl.degree_m f" by (rule sym, rule poly_mod.degree_m_eq[OF _ pl.m1], insert cop p, simp add: l0 p.coprime_exp_mod plp) also have "\ = sum_mset (image_mset pl.degree_m (mset us))" unfolding pl.factorization_m_degree[OF fact plf0] .. also have "\ = sum_list (map pl.degree_m us)" unfolding sum_mset_sum_list[symmetric] by auto also have "\ = sum_list (map degree us)" by (rule arg_cong[OF map_cong, OF refl], rule pl.monic_degree_m, insert mon_gs, auto) finally have degf_gs: "degree f = sum_list (map degree us)" by auto hence gs: "us \ []" using degf by (cases us, auto) from choose_u_member[OF gs] have u_gs: "u \ set us" unfolding u_def by auto from fact u_gs have irred: "pl.irreducible\<^sub>d_m u" unfolding pl.factorization_m_def by auto hence deg_u: "degree u \ 0" unfolding pl.irreducible\<^sub>d_m_def norm[OF u_gs] by auto have deg_uf: "degree u \ degree f" unfolding degf_gs using split_list[OF u_gs] by auto from mon_gs[OF u_gs] have mon_u: "monic u" and u0: "u \ 0" by auto have f0: "f \ 0" using degf by auto from norm have norm': "image_mset pl.Mp (mset us) = mset us" by (induct us, auto) have pl0: "pl \ 0" using pl1 by auto note short_main = LLL_short_polynomial[OF deg_u deg_uf pl1 mon_u] from short_main(1-2)[folded g_def] have "degree k < degree f" unfolding k_def by (smt Suc_leI Suc_less_eq degree_gcd1 gcd.commute le_imp_less_Suc le_trans) hence deg_fk: "(degree k = 0 \ degree f \ degree k) = (degree k = 0)" by auto note res = res[unfolded LLL_reconstruction.simps[of f us] Let_def, folded u_def, folded g_def, folded k_def, unfolded deg_fk] show ?case proof (cases "degree k = 0") case True with res have fs: "fs = [f]" by auto from sf fF have sf: "p.square_free_m f" using p.square_free_m_factor(1)[of f] unfolding dvd_def by auto have irr: "irreducible\<^sub>d f" proof (rule ccontr) assume "\ irreducible\<^sub>d f" from reducible\<^sub>dE[OF this] degf obtain f1 f2 where f: "f = f1 * f2" and deg12: "degree f1 \ 0" "degree f2 \ 0" "degree f1 < degree f" "degree f2 < degree f" by (simp, metis) from pl.unique_factorization_m_factor[OF p uf[unfolded f], folded f, OF cop sf l0 plp] obtain us1 us2 where uf12: "pl.unique_factorization_m f1 (lead_coeff f1, us1)" "pl.unique_factorization_m f2 (lead_coeff f2, us2)" and gs: "mset us = us1 + us2" and norm12: "image_mset pl.Mp us2 = us2" "image_mset pl.Mp us1 = us1" unfolding pl.Mf_def norm' split by (auto simp: pl.Mf_def) note norm_u = norm[OF u_gs] from u_gs have u_gs': "u \# mset us" by auto with pl.factorization_m_mem_dvdm[OF fact, of u] have u_f: "pl.dvdm u f" by auto from u_gs'[unfolded gs] have "u \# us1 \ u \# us2" by auto with pl.factorization_m_mem_dvdm[OF to_fact[OF uf12(1)], of u] pl.factorization_m_mem_dvdm[OF to_fact[OF uf12(2)], of u] have "pl.dvdm u f1 \ pl.dvdm u f2" unfolding norm12 norm_u by auto from this have "\ f1 f2. f = f1 * f2 \ degree f1 \ 0 \ degree f2 \ 0 \ degree f1 < degree f \ degree f2 < degree f \ pl.dvdm u f1" proof assume "pl.dvdm u f1" thus ?thesis using f deg12 by auto next from f have f: "f = f2 * f1" by auto assume "pl.dvdm u f2" thus ?thesis using f deg12 by auto qed then obtain f1 f2 where prod: "f = f1 * f2" and deg: "degree f1 \ 0" "degree f2 \ 0" "degree f1 < degree f" "degree f2 < degree f" and uf1: "pl.dvdm u f1" by auto from pl.unique_factorization_m_factor[OF p uf[unfolded prod], folded prod, OF cop sf l0 plp] obtain us1 where fact_f1: "pl.unique_factorization_m f1 (lead_coeff f1, us1)" by auto have plf1: "pl.Mp f1 \ 0" using to_fact[OF fact_f1] pl.factorization_m_lead_coeff pl.unique_factorization_m_zero fact_f1 by fastforce have "degree u \ degree f1" by (rule pl.dvdm_degree[OF mon_u uf1 plf1]) with deg have deg_uf: "degree u < degree f" by auto have pl0: "pl \ 0" using pl.m1 plp by linarith let ?n = "degree f" let ?n1 = "degree f1" let ?d = "degree u" from prod fF have f1F: "f1 dvd F" unfolding dvd_def by auto from deg_uf have deg_uf': "?d \ ?n" by auto from deg have f1_0: "f1 \ 0" by auto have ug: "pl.dvdm u g" using short_main(3) unfolding g_def . have g0: "g \ 0" using short_main(2) unfolding g_def . have deg_gf: "degree g < degree f" using short_main(1) unfolding g_def . let ?N = "degree F" from fF prod have f1F: "f1 dvd F" unfolding dvd_def by auto have "\g\\<^sup>2 \ 2 ^ (?n - 1) * \f1\\<^sup>2" unfolding g_def by (rule short_main(4)[OF deg_uf _ uf1], insert deg, auto) also have "\ \ 2 ^ (?n - 1) * (2 ^ (2 * degree f1) * \F\\<^sup>2)" by (rule mult_left_mono[OF sq_norm_factor_bound[OF f1F F0]], simp) also have "\ = 2 ^ ((?n - 1) + 2 * degree f1) * \F\\<^sup>2" unfolding power_add by simp also have "\ \ 2 ^ ((?n - 1) + 2 * (?n - 1)) * \F\\<^sup>2" by (rule mult_right_mono, insert deg(3), auto) also have "\ = 2 ^ (3 * (?n - 1)) * \F\\<^sup>2" by simp finally have ineq_g: "\g\\<^sup>2 \ 2 ^ (3 * (?n - 1)) * \F\\<^sup>2" . from power_mono[OF this, of ?n1] have ineq1: "\g\\<^sup>2 ^ ?n1 \ (2 ^ (3 * (?n - 1)) * \F\\<^sup>2)^?n1" by auto from F0 have normF: "\F\\<^sup>2 \ 1" using sq_norm_poly_pos[of F] by presburger from g0 have normg: "\g\\<^sup>2 \ 1" using sq_norm_poly_pos[of g] by presburger from f0 have normf: "\f\\<^sup>2 \ 1" using sq_norm_poly_pos[of f] by presburger from f1_0 have normf1: "\f1\\<^sup>2 \ 1" using sq_norm_poly_pos[of f1] by presburger from power_mono[OF sq_norm_factor_bound[OF f1F F0], of "degree g"] have ineq2: "\f1\\<^sup>2 ^ degree g \ (2 ^ (2 * ?n1) * \F\\<^sup>2) ^ degree g" by auto also have "\ \ (2 ^ (2 * ?n1) * \F\\<^sup>2) ^ (?n - 1)" by (rule pow_mono_exp, insert deg_gf normF, auto) finally have ineq2: "\f1\\<^sup>2 ^ degree g \ (2 ^ (2 * ?n1) * \F\\<^sup>2) ^ (?n - 1)" . have nN: "?n \ ?N" using fF F0 by (metis dvd_imp_degree_le) from deg nN have n1N: "?n1 \ ?N - 1" by auto have "\f1\\<^sup>2 ^ degree g * \g\\<^sup>2 ^ ?n1 \ (2 ^ (2 * ?n1) * \F\\<^sup>2) ^ (?n - 1) * (2 ^ (3 * (?n - 1)) * \F\\<^sup>2)^?n1" by (rule mult_mono[OF ineq2 ineq1], force+) also have "\ \ (2 ^ (2 * (?N - 1)) * \F\\<^sup>2) ^ (?N - 1) * (2 ^ (3 * (?N - 1)) * \F\\<^sup>2) ^ (?N - 1)" by (rule mult_mono[OF power_both_mono[OF _ _ mult_mono] power_both_mono], insert normF n1N nN, auto intro: power_both_mono mult_mono) also have "\ = 2 ^ (2 * (?N -1) * (?N - 1) + 3 * (?N - 1) * (?N - 1)) * (\F\\<^sup>2)^((?N - 1) + (?N - 1))" unfolding power_mult_distrib power_add power_mult by simp also have "2 * (?N - 1) * (?N - 1) + 3 * (?N - 1) * (?N - 1) = 5 * (?N - 1) * (?N - 1)" by simp also have "?N - 1 + (?N - 1) = 2 * (?N - 1)" by simp also have "2^(5 * (?N - 1) * (?N - 1)) * \F\\<^sup>2^(2 * (?N - 1)) < pl^2" by (rule large) finally have large: "\f1\\<^sup>2 ^ degree g * \g\\<^sup>2 ^ degree f1 < pl\<^sup>2" . have deg_ug: "degree u \ degree g" proof (rule pl.dvdm_degree[OF mon_u ug], standard) assume "pl.Mp g = 0" from arg_cong[OF this, of "\ p. coeff p (degree g)"] have "pl.M (coeff g (degree g)) = 0" by (auto simp: pl.Mp_def coeff_map_poly) from this[unfolded pl.M_def] obtain c where lg: "lead_coeff g = pl * c" by auto with g0 have c0: "c \ 0" by auto hence "pl^2 \ (lead_coeff g)^2" unfolding lg abs_le_square_iff[symmetric] by (rule aux_abs_int) also have "\ \ \g\\<^sup>2 ^ 1" using coeff_le_sq_norm[of g] by auto also have "\ \ \g\\<^sup>2 ^ degree f1" by (rule pow_mono_exp, insert deg normg, auto) also have "\ = 1 * \" by simp also have "\ \ \f1\\<^sup>2 ^ degree g * \g\\<^sup>2 ^ degree f1" by (rule mult_right_mono, insert normf1, auto) also have "\ < pl\<^sup>2" by (rule large) finally show False by auto qed from deg deg_u deg_ug have "degree f1 > 0" "degree g > 0" by auto from common_factor_via_short[OF this mon_u _ uf1 ug large] deg_u pl.m1 have "0 < degree (gcd f1 g)" by auto moreover from True[unfolded k_def] have "degree (gcd f g) = 0" . moreover have dvd: "gcd f1 g dvd gcd f g" using f0 unfolding prod by simp ultimately show False using divides_degree[OF dvd] using f0 by simp qed show ?thesis unfolding fs using irr by auto next case False define f1 where "f1 = f div k" have f: "f = f1 * k" unfolding f1_def k_def by auto with arg_cong[OF this, of degree] f0 have deg_f1k: "degree f = degree f1 + degree k" by (auto simp: degree_mult_eq) from f fF have dvd: "f1 dvd F" "k dvd F" unfolding dvd_def by auto obtain gs1 gs2 where part: "List.partition (\gi. p.dvdm gi f1) us = (gs1, gs2)" by force note IH = 1(1-2)[OF refl u_def g_def k_def refl, unfolded deg_fk, OF False f1_def part[symmetric] refl] obtain fs1 where fs1: "LLL_reconstruction f1 gs1 = fs1" by auto obtain fs2 where fs2: "LLL_reconstruction k gs2 = fs2" by auto from False res[folded f1_def, unfolded part split fs1 fs2] have fs: "fs = fs1 @ fs2" by auto from short_main(1) have deg_gf: "degree g < degree f" unfolding g_def by auto from short_main(2) have g0: "g \ 0" unfolding g_def by auto have deg_kg: "degree k \ degree g" unfolding k_def gcd.commute[of f g] by (rule degree_gcd1[OF g0]) from deg_gf deg_kg have deg_kf: "degree k < degree f" by auto with deg_f1k have deg_f1: "degree f1 \ 0" by auto have sf_f: "p.square_free_m f" using sf fF p.square_free_m_factor unfolding dvd_def by blast from p.unique_factorization_m_factor_partition[OF l0 uf[unfolded plp] f cop sf_f part] have uf: "pl.unique_factorization_m f1 (lead_coeff f1, mset gs1)" "pl.unique_factorization_m k (lead_coeff k, mset gs2)" by (auto simp: plp) have "set us = set gs1 \ set gs2" using part by auto with norm have norm_12: "gi \ set gs1 \ gi \ set gs2 \ pl.Mp gi = gi" for gi by auto note IH1 = IH(1)[OF fs1 deg_f1 uf(1) dvd(1) norm_12] note IH2 = IH(2)[OF fs2 False uf(2) dvd(2) norm_12] show ?thesis unfolding fs f using IH1 IH2 by auto qed qed qed lemma LLL_many_reconstruction: assumes "LLL_many_reconstruction f us = fs" and "degree f \ 0" and "poly_mod.unique_factorization_m pl f (lead_coeff f, mset us)" and "f dvd F" and "\ ui. ui \ set us \ poly_mod.Mp pl ui = ui" and F0: "F \ 0" and cop: "coprime (lead_coeff F) p" and sf: "poly_mod.square_free_m p F" and pl1: "pl > 1" and plp: "pl = p^l" and p: "prime p" and large: "2^(5 * (degree F div 2) * (degree F div 2)) * \F\\<^sup>2^(2 * (degree F div 2)) < pl\<^sup>2" shows "f = prod_list fs \ (\ fi \ set fs. irreducible\<^sub>d fi)" proof - interpret p: poly_mod_prime p by (standard, rule p) interpret pl: poly_mod_2 "pl" by (standard, rule pl1) from pl1 plp have l0: "l \ 0" by (cases l, auto) show ?thesis using assms(1-5) proof (induct f us arbitrary: fs rule: LLL_many_reconstruction.induct) case (1 f us fs) note res = 1(3) note degf = 1(4) note uf = 1(5) note fF = 1(6) note norm = 1(7) note to_fact = pl.unique_factorization_m_imp_factorization note fact = to_fact[OF uf] have mon_gs: "ui \ set us \ monic ui" for ui using norm fact unfolding pl.factorization_m_def by auto from p.coprime_lead_coeff_factor[OF p.prime] fF cop have cop: "coprime (lead_coeff f) p" unfolding dvd_def by blast have plf0: "pl.Mp f \ 0" using fact pl.factorization_m_lead_coeff pl.unique_factorization_m_zero uf by fastforce have "degree f = pl.degree_m f" by (rule sym, rule poly_mod.degree_m_eq[OF _ pl.m1], insert cop p, simp add: l0 p.coprime_exp_mod plp) also have "\ = sum_mset (image_mset pl.degree_m (mset us))" unfolding pl.factorization_m_degree[OF fact plf0] .. also have "\ = sum_list (map pl.degree_m us)" unfolding sum_mset_sum_list[symmetric] by auto also have "\ = sum_list (map degree us)" by (rule arg_cong[OF map_cong, OF refl], rule pl.monic_degree_m, insert mon_gs, auto) finally have degf_gs: "degree f = sum_list (map degree us)" by auto hence gs: "us \ []" using degf by (cases us, auto) from 1(4) have f0: "f \ 0" and df0: "degree f \ 0" by auto from norm have norm': "image_mset pl.Mp (mset us) = mset us" by (induct us, auto) have pl0: "pl \ 0" using pl1 by auto let ?D2 = "degree F div 2" let ?d2 = "degree f div 2" define gg where "gg = LLL_short_polynomial pl (Suc ?d2)" let ?us = "filter (\u. degree u \ ?d2) us" note res = res[unfolded LLL_many_reconstruction.simps[of f us], unfolded Let_def, folded gg_def] let ?f2_opt = "find_map_filter (\u. gcd f (gg u)) (\f2. 0 < degree f2 \ degree f2 < degree f) ?us" show ?case proof (cases ?f2_opt) case (Some f2) from find_map_filter_Some[OF this] obtain g where deg_f2: "degree f2 \ 0" "degree f2 < degree f" and dvd: "f2 dvd f" and gcd: "f2 = gcd f g" by auto note res = res[unfolded Some option.simps] define f1 where "f1 = f div f2" have f: "f = f1 * f2" unfolding f1_def using dvd by auto with arg_cong[OF this, of degree] f0 have deg_sum: "degree f = degree f1 + degree f2" by (auto simp: degree_mult_eq) with deg_f2 have deg_f1: "degree f1 \ 0" "degree f1 < degree f" by auto from f fF have dvd: "f1 dvd F" "f2 dvd F" unfolding dvd_def by auto obtain gs1 gs2 where part: "List.partition (\gi. p.dvdm gi f1) us = (gs1, gs2)" by force note IH = 1(1-2)[OF refl refl refl, unfolded Let_def, folded gg_def, OF Some f1_def part[symmetric] refl] obtain fs1 where fs1: "LLL_many_reconstruction f1 gs1 = fs1" by blast obtain fs2 where fs2: "LLL_many_reconstruction f2 gs2 = fs2" by blast from res[folded f1_def, unfolded part split fs1 fs2] have fs: "fs = fs1 @ fs2" by auto have sf_f: "p.square_free_m f" using sf fF p.square_free_m_factor unfolding dvd_def by blast from p.unique_factorization_m_factor_partition[OF l0 uf[unfolded plp] f cop sf_f part] have uf: "pl.unique_factorization_m f1 (lead_coeff f1, mset gs1)" "pl.unique_factorization_m f2 (lead_coeff f2, mset gs2)" by (auto simp: plp) have "set us = set gs1 \ set gs2" using part by auto with norm have norm_12: "gi \ set gs1 \ gi \ set gs2 \ pl.Mp gi = gi" for gi by auto note IH1 = IH(1)[OF fs1 deg_f1(1) uf(1) dvd(1) norm_12] note IH2 = IH(2)[OF fs2 deg_f2(1) uf(2) dvd(2) norm_12] show ?thesis unfolding fs f using IH1 IH2 by auto next case None from res[unfolded None option.simps] have fs_f: "fs = [f]" by simp from sf fF have sf: "p.square_free_m f" using p.square_free_m_factor(1)[of f] unfolding dvd_def by auto have "irreducible\<^sub>d f" proof (rule ccontr) assume "\ irreducible\<^sub>d f" from reducible\<^sub>dE[OF this] degf obtain f1 f2 where f: "f = f1 * f2" and deg12: "degree f1 \ 0" "degree f2 \ 0" "degree f1 < degree f" "degree f2 < degree f" by (simp, metis) from f0 have "degree f = degree f1 + degree f2" unfolding f by (auto simp: degree_mult_eq) hence "degree f1 \ degree f div 2 \ degree f2 \ degree f div 2" by auto then obtain f1 f2 where f: "f = f1 * f2" and deg12: "degree f1 \ 0" "degree f2 \ 0" "degree f1 \ degree f div 2" "degree f2 < degree f" proof (standard, goal_cases) case 1 from 1(1)[of f1 f2] 1(2) f deg12 show ?thesis by auto next case 2 from 2(1)[of f2 f1] 2(2) f deg12 show ?thesis by auto qed from f0 f have f10: "f1 \ 0" by auto from sf f have sf1: "p.square_free_m f1" using p.square_free_m_factor(1)[of f1] by auto from p.coprime_lead_coeff_factor[OF p.prime cop[unfolded f]] have cop1: "coprime (lead_coeff f1) p" by auto have deg_m1: "pl.degree_m f1 = degree f1" by (rule poly_mod.degree_m_eq[OF _ pl.m1], insert cop1 p, simp add: l0 p.coprime_exp_mod plp) from pl.unique_factorization_m_factor[OF p uf[unfolded f], folded f, OF cop sf l0 plp] obtain us1 us2 where uf12: "pl.unique_factorization_m f1 (lead_coeff f1, us1)" "pl.unique_factorization_m f2 (lead_coeff f2, us2)" and gs: "mset us = us1 + us2" and norm12: "image_mset pl.Mp us2 = us2" "image_mset pl.Mp us1 = us1" unfolding pl.Mf_def norm' split by (auto simp: pl.Mf_def) from gs have "x \# us1 \ x \# mset us" for x by auto hence sub1: "x \# us1 \ x \ set us" for x by auto from to_fact[OF uf12(1)] have fact1: "pl.factorization_m f1 (lead_coeff f1, us1)" . have plf10: "pl.Mp f1 \ 0" using fact1 pl.factorization_m_lead_coeff pl.unique_factorization_m_zero uf12(1) by fastforce have "degree f1 = pl.degree_m f1" using deg_m1 by simp also have "\ = sum_mset (image_mset pl.degree_m us1)" unfolding pl.factorization_m_degree[OF fact1 plf10] .. also have "\ = sum_mset (image_mset degree us1)" by (rule arg_cong[of _ _ sum_mset], rule image_mset_cong, rule pl.monic_degree_m, rule mon_gs, rule sub1) finally have degf1_sum: "degree f1 = sum_mset (image_mset degree us1)" by auto with deg12 have "us1 \ {#}" by auto then obtain u us11 where us1: "us1 = {#u#} + us11" by (cases us1, auto) hence u1: "u \# us1" by auto hence u: "u \ set us" by (rule sub1) let ?g = "gg u" from pl.factorization_m_mem_dvdm[OF fact1, of u] u1 have u_f1: "pl.dvdm u f1" by auto note norm_u = norm[OF u] from fact u have irred: "pl.irreducible\<^sub>d_m u" unfolding pl.factorization_m_def by auto hence deg_u: "degree u \ 0" unfolding pl.irreducible\<^sub>d_m_def norm[OF u] by auto have "degree u \ degree f1" unfolding degf1_sum unfolding us1 by simp also have "\ \ degree f div 2" by fact finally have deg_uf: "degree u \ degree f div 2" . hence deg_uf': "degree u \ Suc (degree f div 2)" "degree u < Suc (degree f div 2)" by auto from mon_gs[OF u] have mon_u: "monic u" . note short = LLL_short_polynomial[OF deg_u deg_uf'(1) pl1 mon_u, folded gg_def] note short = short(1-3) short(4)[OF deg_uf'(2)] from short(1,2) deg12(1,3) f10 have "degree (gcd f ?g) \ degree f div 2" by (metis Suc_leI Suc_le_mono degree_gcd1 gcd.commute le_trans) also have "\ < degree f" using degf by simp finally have "degree (gcd f ?g) < degree f" by simp with find_map_filter_None[OF None, simplified, rule_format, of u] deg_uf u have deg_gcd: "degree (gcd f (?g)) = 0" by (auto simp: gcd.commute) have "gcd f1 (?g) dvd gcd f (?g)" using f0 unfolding f by simp from divides_degree[OF this, unfolded deg_gcd] f0 have deg_gcd1: "degree (gcd f1 (?g)) = 0" by auto from F0 have normF: "\F\\<^sup>2 \ 1" using sq_norm_poly_pos[of F] by presburger have g0: "?g \ 0" using short(2) . from g0 have normg: "\?g\\<^sup>2 \ 1" using sq_norm_poly_pos[of "?g"] by presburger from f10 have normf1: "\f1\\<^sup>2 \ 1" using sq_norm_poly_pos[of f1] by presburger from fF f have f1F: "f1 dvd F" unfolding dvd_def by auto have pl_ge0: "pl \ 0" using pl.poly_mod_2_axioms poly_mod_2_def by auto from fF have "degree f \ degree F" using F0 f0 by (metis dvd_imp_degree_le) hence d2D2: "?d2 \ ?D2" by simp with deg12(3) have df1_D2: "degree f1 \ ?D2" by linarith from short(1) d2D2 have dg_D2: "degree (gg u) \ ?D2" by linarith have "\f1\\<^sup>2 ^ degree (gg u) * \gg u\\<^sup>2 ^ degree f1 \ \f1\\<^sup>2 ^ ?D2 * \gg u\\<^sup>2 ^ ?D2" by (rule mult_mono[OF pow_mono_exp pow_mono_exp], insert normf1 normg, auto intro: df1_D2 dg_D2) also have "\ = (\f1\\<^sup>2 * \gg u\\<^sup>2)^?D2" by (simp add: power_mult_distrib) also have "\ \ (\f1\\<^sup>2 * (2^?D2 * \f1\\<^sup>2))^?D2" by (rule power_mono[OF mult_left_mono[OF order.trans[OF short(4)[OF f10 u_f1]]]], insert deg12 d2D2, auto intro!: mult_mono) also have "\ = \f1\\<^sup>2 ^ (?D2 + ?D2) * 2^(?D2 * ?D2)" unfolding power_add power_mult_distrib power_mult by simp also have "\ \ (2 ^ (2 * ?D2) * \F\\<^sup>2) ^ (?D2 + ?D2) * 2^(?D2 * ?D2)" by (rule mult_right_mono[OF order.trans[OF power_mono[OF sq_norm_factor_bound[OF f1F F0]]]], auto intro!: power_mono mult_right_mono df1_D2) also have "\ = 2 ^ (2 * ?D2 * (?D2 + ?D2) + ?D2 * ?D2) * \F\\<^sup>2 ^ (?D2 + ?D2)" unfolding power_mult_distrib power_mult power_add by simp also have "2 * ?D2 * (?D2 + ?D2) + ?D2 * ?D2 = 5 * ?D2 * ?D2" by simp also have "?D2 + ?D2 = 2 * ?D2" by simp finally have large: "\f1\\<^sup>2 ^ degree (gg u) * \gg u\\<^sup>2 ^ degree f1 < pl^2" using large by simp have "degree u \ degree (?g)" proof (rule pl.dvdm_degree[OF mon_u short(3)], standard) assume "pl.Mp (?g) = 0" from arg_cong[OF this, of "\ p. coeff p (degree ?g)"] have "pl.M (coeff ?g (degree ?g)) = 0" by (auto simp: pl.Mp_def coeff_map_poly) from this[unfolded pl.M_def] obtain c where lg: "lead_coeff ?g = pl * c" by auto with g0 have c0: "c \ 0" by auto hence "pl^2 \ (lead_coeff ?g)^2" unfolding lg abs_le_square_iff[symmetric] by (rule aux_abs_int) also have "\ \ \?g\\<^sup>2" using coeff_le_sq_norm[of ?g] by auto also have "\ = \?g\\<^sup>2 ^ 1" by simp also have "\ \ \?g\\<^sup>2 ^ degree f1" by (rule pow_mono_exp, insert deg12 normg, auto) also have "\ = 1 * \" by simp also have "\ \ \f1\\<^sup>2 ^ degree ?g * \?g\\<^sup>2 ^ degree f1" by (rule mult_right_mono, insert normf1, auto) also have "\ < pl\<^sup>2" by (rule large) finally show False by auto qed with deg_u have deg_g: "0 < degree (gg u)" by auto have pl_ge0: "pl \ 0" using pl.poly_mod_2_axioms poly_mod_2_def by auto from fF have "degree f \ degree F" using F0 f0 by (metis dvd_imp_degree_le) hence d2D2: "?d2 \ ?D2" by simp with deg12(3) have df1_D2: "degree f1 \ ?D2" by linarith from short(1) d2D2 have dg_D2: "degree (gg u) \ ?D2" by linarith have "0 < degree f1" "0 < degree u" using deg12 deg_u by auto from common_factor_via_short[of f1 "gg u", OF this(1) deg_g mon_u this(2) u_f1 short(3) _ pl_ge0] deg_gcd1 have "pl^2 \ \f1\\<^sup>2 ^ degree (gg u) * \gg u\\<^sup>2 ^ degree f1" by linarith also have "\ < pl^2" by (rule large) finally show False by simp qed thus ?thesis using fs_f by simp qed qed qed end lemma LLL_factorization: assumes res: "LLL_factorization f = gs" and sff: "square_free f" and deg: "degree f \ 0" shows "f = prod_list gs \ (\g\set gs. irreducible\<^sub>d g)" proof - let ?lc = "lead_coeff f" define p where "p \ suitable_prime_bz f" obtain c gs where fff: "finite_field_factorization_int p f = (c,gs)" by force let ?degs = "map degree gs" note res = res[unfolded LLL_factorization_def Let_def, folded p_def, unfolded fff split, folded] from suitable_prime_bz[OF sff refl] have prime: "prime p" and cop: "coprime ?lc p" and sf: "poly_mod.square_free_m p f" unfolding p_def by auto note res from prime interpret p: poly_mod_prime p by unfold_locales define K where "K = 2^(5 * (degree f - 1) * (degree f - 1)) * \f\\<^sup>2^(2 * (degree f - 1))" define N where "N = sqrt_int_ceiling K" have K0: "K \ 0" unfolding K_def by fastforce have N0: "N \ 0" unfolding N_def sqrt_int_ceiling using K0 by (smt of_int_nonneg real_sqrt_ge_0_iff zero_le_ceiling) define n where "n = find_exponent p N" note res = res[folded n_def[unfolded N_def K_def]] note n = find_exponent[OF p.m1, of N, folded n_def] note bh = p.berlekamp_and_hensel_separated(1)[OF cop sf refl fff n(2)] from deg have f0: "f \ 0" by auto from n p.m1 have pn1: "p ^ n > 1" by auto note res = res[folded bh(1)] note * = p.berlekamp_hensel_unique[OF cop sf bh n(2)] note ** = p.berlekamp_hensel_main[OF n(2) bh cop sf fff] from res * ** have uf: "poly_mod.unique_factorization_m (p ^ n) f (lead_coeff f, mset (berlekamp_hensel p n f))" and norm: "\ui. ui \ set (berlekamp_hensel p n f) \ poly_mod.Mp (p ^ n) ui = ui" unfolding berlekamp_hensel_def fff split by auto have K: "K < (p ^ n)\<^sup>2" using n sqrt_int_ceiling_bound[OF K0] by (smt N0 N_def n(1) power2_le_imp_le) show ?thesis by (rule LLL_implementation.LLL_reconstruction[OF res deg uf dvd_refl norm f0 cop sf pn1 refl prime K[unfolded K_def]]) qed lemma LLL_many_factorization: assumes res: "LLL_many_factorization f = gs" and sff: "square_free f" and deg: "degree f \ 0" shows "f = prod_list gs \ (\g\set gs. irreducible\<^sub>d g)" proof - let ?lc = "lead_coeff f" define p where "p \ suitable_prime_bz f" obtain c gs where fff: "finite_field_factorization_int p f = (c,gs)" by force let ?degs = "map degree gs" note res = res[unfolded LLL_many_factorization_def Let_def, folded p_def, unfolded fff split, folded] from suitable_prime_bz[OF sff refl] have prime: "prime p" and cop: "coprime ?lc p" and sf: "poly_mod.square_free_m p f" unfolding p_def by auto note res from prime interpret p: poly_mod_prime p by unfold_locales define K where "K = 2^(5 * (degree f div 2) * (degree f div 2)) * \f\\<^sup>2^(2 * (degree f div 2))" define N where "N = sqrt_int_ceiling K" have K0: "K \ 0" unfolding K_def by fastforce have N0: "N \ 0" unfolding N_def sqrt_int_ceiling using K0 by (smt of_int_nonneg real_sqrt_ge_0_iff zero_le_ceiling) define n where "n = find_exponent p N" note res = res[folded n_def[unfolded N_def K_def]] note n = find_exponent[OF p.m1, of N, folded n_def] note bh = p.berlekamp_and_hensel_separated(1)[OF cop sf refl fff n(2)] from deg have f0: "f \ 0" by auto from n p.m1 have pn1: "p ^ n > 1" by auto note res = res[folded bh(1)] note * = p.berlekamp_hensel_unique[OF cop sf bh n(2)] note ** = p.berlekamp_hensel_main[OF n(2) bh cop sf fff] from res * ** have uf: "poly_mod.unique_factorization_m (p ^ n) f (lead_coeff f, mset (berlekamp_hensel p n f))" and norm: "\ui. ui \ set (berlekamp_hensel p n f) \ poly_mod.Mp (p ^ n) ui = ui" unfolding berlekamp_hensel_def fff split by auto have K: "K < (p ^ n)\<^sup>2" using n sqrt_int_ceiling_bound[OF K0] by (smt N0 N_def n(1) power2_le_imp_le) show ?thesis by (rule LLL_implementation.LLL_many_reconstruction[OF res deg uf dvd_refl norm f0 cop sf pn1 refl prime K[unfolded K_def]]) qed lift_definition one_lattice_LLL_factorization :: int_poly_factorization_algorithm is LLL_factorization using LLL_factorization by auto lift_definition many_lattice_LLL_factorization :: int_poly_factorization_algorithm is LLL_many_factorization using LLL_many_factorization by auto lemma LLL_factorization_primitive: assumes "LLL_factorization f = fs" "square_free f" "0 < degree f" "primitive f" shows "f = prod_list fs \ (\fi\set fs. irreducible fi \ 0 < degree fi \ primitive fi)" using assms(1) by (intro int_poly_factorization_algorithm_irreducible[of one_lattice_LLL_factorization, OF _ assms(2-)], transfer, auto) thm factorize_int_poly[of one_lattice_LLL_factorization] thm factorize_int_poly[of many_lattice_LLL_factorization] end diff --git a/thys/Ordinary_Differential_Equations/IVP/Flow.thy b/thys/Ordinary_Differential_Equations/IVP/Flow.thy --- a/thys/Ordinary_Differential_Equations/IVP/Flow.thy +++ b/thys/Ordinary_Differential_Equations/IVP/Flow.thy @@ -1,3199 +1,3194 @@ section \Flow\ theory Flow imports Picard_Lindeloef_Qualitative "HOL-Library.Diagonal_Subsequence" "../Library/Bounded_Linear_Operator" "../Library/Multivariate_Taylor" "../Library/Interval_Integral_HK" begin text \TODO: extend theorems for dependence on initial time\ subsection \simp rules for integrability (TODO: move)\ lemma blinfun_ext: "x = y \ (\i. blinfun_apply x i = blinfun_apply y i)" by transfer auto notation id_blinfun ("1\<^sub>L") lemma blinfun_inverse_left: fixes f::"'a::euclidean_space \\<^sub>L 'a" and f' shows "f o\<^sub>L f' = 1\<^sub>L \ f' o\<^sub>L f = 1\<^sub>L" by transfer (auto dest!: bounded_linear.linear simp: id_def[symmetric] linear_inverse_left) lemma onorm_zero_blinfun[simp]: "onorm (blinfun_apply 0) = 0" by transfer (simp add: onorm_zero) lemma blinfun_compose_1_left[simp]: "x o\<^sub>L 1\<^sub>L = x" and blinfun_compose_1_right[simp]: "1\<^sub>L o\<^sub>L y = y" by (auto intro!: blinfun_eqI) named_theorems integrable_on_simps lemma integrable_on_refl_ivl[intro, simp]: "g integrable_on {b .. (b::'b::ordered_euclidean_space)}" and integrable_on_refl_closed_segment[intro, simp]: "h integrable_on closed_segment a a" using integrable_on_refl by auto lemma integrable_const_ivl_closed_segment[intro, simp]: "(\x. c) integrable_on closed_segment a (b::real)" by (auto simp: closed_segment_eq_real_ivl) lemma integrable_ident_ivl[intro, simp]: "(\x. x) integrable_on closed_segment a (b::real)" and integrable_ident_cbox[intro, simp]: "(\x. x) integrable_on cbox a (b::real)" by (auto simp: closed_segment_eq_real_ivl ident_integrable_on) lemma content_closed_segment_real: fixes a b::real shows "content (closed_segment a b) = abs (b - a)" by (auto simp: closed_segment_eq_real_ivl) lemma integral_const_closed_segment: fixes a b::real shows "integral (closed_segment a b) (\x. c) = abs (b - a) *\<^sub>R c" by (auto simp: closed_segment_eq_real_ivl content_closed_segment_real) lemmas [integrable_on_simps] = integrable_on_empty \ \empty\ integrable_on_refl integrable_on_refl_ivl integrable_on_refl_closed_segment \ \singleton\ integrable_const integrable_const_ivl integrable_const_ivl_closed_segment \ \constant\ ident_integrable_on integrable_ident_ivl integrable_ident_cbox \ \identity\ lemma integrable_cmul_real: fixes K::real shows "f integrable_on X \ (\x. K * f x) integrable_on X " unfolding real_scaleR_def[symmetric] by (rule integrable_cmul) lemmas [integrable_on_simps] = integrable_0 integrable_neg integrable_cmul integrable_cmul_real integrable_on_cmult_iff integrable_on_cmult_left integrable_on_cmult_right integrable_on_cdivide integrable_on_cmult_iff integrable_on_cmult_left_iff integrable_on_cmult_right_iff integrable_on_cdivide_iff integrable_diff integrable_add integrable_sum lemma dist_cancel_add1: "dist (t0 + et) t0 = norm et" by (simp add: dist_norm) lemma double_nonneg_le: fixes a::real shows "a * 2 \ b \ a \ 0 \ a \ b" by arith subsection \Nonautonomous IVP on maximal existence interval\ context ll_on_open_it begin context fixes x0 assumes iv_defined: "t0 \ T" "x0 \ X" begin lemmas closed_segment_iv_subset_domain = closed_segment_subset_domainI[OF iv_defined(1)] lemma local_unique_solutions: obtains t u L where "0 < t" "0 < u" "cball t0 t \ existence_ivl t0 x0" "cball x0 (2 * u) \ X" "\t'. t' \ cball t0 t \ L-lipschitz_on (cball x0 (2 * u)) (f t')" "\x. x \ cball x0 u \ (flow t0 x usolves_ode f from t0) (cball t0 t) (cball x u)" "\x. x \ cball x0 u \ cball x u \ X" proof - from local_unique_solution[OF iv_defined] obtain et ex B L where "0 < et" "0 < ex" "cball t0 et \ T" "cball x0 ex \ X" "unique_on_cylinder t0 (cball t0 et) x0 ex f B L" by metis then interpret cyl: unique_on_cylinder t0 "cball t0 et" x0 ex "cball x0 ex" f B L by auto from cyl.solution_solves_ode order_refl \cball x0 ex \ X\ have "(cyl.solution solves_ode f) (cball t0 et) X" by (rule solves_ode_on_subset) then have "cball t0 et \ existence_ivl t0 x0" by (rule existence_ivl_maximal_interval) (insert \cball t0 et \ T\ \0 < et\, auto) have "cball t0 et = {t0 - et .. t0 + et}" using \et > 0\ by (auto simp: dist_real_def) then have cylbounds[simp]: "cyl.tmin = t0 - et" "cyl.tmax = t0 + et" unfolding cyl.tmin_def cyl.tmax_def using \0 < et\ by auto define et' where "et' \ et / 2" define ex' where "ex' \ ex / 2" have "et' > 0" "ex' > 0" using \0 < et\ \0 < ex\ by (auto simp: et'_def ex'_def) moreover from \cball t0 et \ existence_ivl t0 x0\ have "cball t0 et' \ existence_ivl t0 x0" by (force simp: et'_def dest!: double_nonneg_le) moreover from this have "cball t0 et' \ T" using existence_ivl_subset[of x0] by simp have "cball x0 (2 * ex') \ X" "\t'. t' \ cball t0 et' \ L-lipschitz_on (cball x0 (2 * ex')) (f t')" using cyl.lipschitz \0 < et\ \cball x0 ex \ X\ by (auto simp: ex'_def et'_def intro!:) moreover { fix x0'::'a assume x0': "x0' \ cball x0 ex'" { fix b assume d: "dist x0' b \ ex'" have "dist x0 b \ dist x0 x0' + dist x0' b" by (rule dist_triangle) also have "\ \ ex' + ex'" using x0' d by simp also have "\ \ ex" by (simp add: ex'_def) finally have "dist x0 b \ ex" . } note triangle = this have subs1: "cball t0 et' \ cball t0 et" and subs2: "cball x0' ex' \ cball x0 ex" and subs: "cball t0 et' \ cball x0' ex' \ cball t0 et \ cball x0 ex" using \0 < ex\ \0 < et\ x0' by (auto simp: ex'_def et'_def triangle dest!: double_nonneg_le) have subset_X: "cball x0' ex' \ X" using \cball x0 ex \ X\ subs2 \0 < ex'\ by force then have "x0' \ X" using \0 < ex'\ by force have x0': "t0 \ T" "x0' \ X" by fact+ have half_intros: "a \ ex' \ a \ ex" "a \ et' \ a \ et" and halfdiv_intro: "a * 2 \ ex / B \ a \ ex' / B" for a using \0 < ex\ \0 < et\ by (auto simp: ex'_def et'_def) interpret cyl': solution_in_cylinder t0 "cball t0 et'" x0' ex' f "cball x0' ex'" B using \0 < et'\ \0 < ex'\ \0 < et\ cyl.norm_f cyl.continuous subs1 \cball t0 et \ T\ apply unfold_locales apply (auto simp: split_beta' dist_cancel_add1 intro!: triangle continuous_intros cyl.norm_f order_trans[OF _ cyl.e_bounded] halfdiv_intro) by (simp add: ex'_def et'_def dist_commute) interpret cyl': unique_on_cylinder t0 "cball t0 et'" x0' ex' "cball x0' ex'" f B L using cyl.lipschitz[simplified] subs subs1 by (unfold_locales) (auto simp: triangle intro!: half_intros lipschitz_on_subset[OF _ subs2]) from cyl'.solution_usolves_ode have "(flow t0 x0' usolves_ode f from t0) (cball t0 et') (cball x0' ex')" apply (rule usolves_ode_solves_odeI) subgoal apply (rule cyl'.solves_ode_on_subset_domain[where Y=X]) subgoal apply (rule solves_ode_on_subset[where S="existence_ivl t0 x0'" and Y=X]) subgoal by (rule flow_solves_ode[OF x0']) subgoal using subs2 \cball x0 ex \ X\ \0 < et'\ \cball t0 et' \ T\ by (intro existence_ivl_maximal_interval[OF solves_ode_on_subset[OF cyl'.solution_solves_ode]]) auto subgoal by force done subgoal by (force simp: \x0' \ X\ iv_defined) subgoal using \0 < et'\ by force subgoal by force subgoal by force done subgoal by (force simp: \x0' \ X\ iv_defined cyl'.solution_iv) done note this subset_X } ultimately show thesis .. qed lemma Picard_iterate_mem_existence_ivlI: assumes "t \ T" assumes "compact C" "x0 \ C" "C \ X" assumes "\y s. s \ {t0 -- t} \ y t0 = x0 \ y \ {t0--s} \ C \ continuous_on {t0--s} y \ x0 + ivl_integral t0 s (\t. f t (y t)) \ C" shows "t \ existence_ivl t0 x0" "\s. s \ {t0 -- t} \ flow t0 x0 s \ C" proof - have "{t0 -- t} \ T" by (intro closed_segment_subset_domain iv_defined assms) from lipschitz_on_compact[OF compact_segment \{t0 -- t} \ T\ \compact C\ \C \ X\] obtain L where L: "\s. s \ {t0 -- t} \ L-lipschitz_on C (f s)" by metis interpret uc: unique_on_closed t0 "{t0 -- t}" x0 f C L using assms closed_segment_iv_subset_domain by unfold_locales (auto intro!: L compact_imp_closed \compact C\ continuous_on_f continuous_intros simp: split_beta) have "{t0 -- t} \ existence_ivl t0 x0" using assms closed_segment_iv_subset_domain by (intro maximal_existence_flow[OF solves_ode_on_subset[OF uc.solution_solves_ode]]) auto thus "t \ existence_ivl t0 x0" using assms by auto show "flow t0 x0 s \ C" if "s \ {t0 -- t}" for s proof - have "flow t0 x0 s = uc.solution s" "uc.solution s \ C" using solves_odeD[OF uc.solution_solves_ode] that assms by (auto simp: closed_segment_iv_subset_domain intro!: maximal_existence_flowI(2)[where K="{t0 -- t}"]) thus ?thesis by simp qed qed lemma flow_has_vderiv_on: "(flow t0 x0 has_vderiv_on (\t. f t (flow t0 x0 t))) (existence_ivl t0 x0)" by (rule solves_ode_vderivD[OF flow_solves_ode[OF iv_defined]]) lemmas flow_has_vderiv_on_compose[derivative_intros] = has_vderiv_on_compose2[OF flow_has_vderiv_on, THEN has_vderiv_on_eq_rhs] end lemma unique_on_intersection: assumes sols: "(x solves_ode f) U X" "(y solves_ode f) V X" assumes iv_mem: "t0 \ U" "t0 \ V" and subs: "U \ T" "V \ T" assumes ivls: "is_interval U" "is_interval V" assumes iv: "x t0 = y t0" assumes mem: "t \ U" "t \ V" shows "x t = y t" proof - from maximal_existence_flow(2)[OF sols(1) refl ivls(1) iv_mem(1) subs(1) mem(1)] maximal_existence_flow(2)[OF sols(2) iv[symmetric] ivls(2) iv_mem(2) subs(2) mem(2)] show ?thesis by simp qed lemma unique_solution: assumes sols: "(x solves_ode f) U X" "(y solves_ode f) U X" assumes iv_mem: "t0 \ U" and subs: "U \ T" assumes ivls: "is_interval U" assumes iv: "x t0 = y t0" assumes mem: "t \ U" shows "x t = y t" by (metis unique_on_intersection assms) lemma assumes s: "s \ existence_ivl t0 x0" assumes t: "t + s \ existence_ivl s (flow t0 x0 s)" shows flow_trans: "flow t0 x0 (s + t) = flow s (flow t0 x0 s) (s + t)" and existence_ivl_trans: "s + t \ existence_ivl t0 x0" proof - note ll_on_open_it_axioms moreover from ll_on_open_it_axioms have iv_defined: "t0 \ T" "x0 \ X" and iv_defined': "s \ T" "flow t0 x0 s \ X" using ll_on_open_it.mem_existence_ivl_iv_defined s t by blast+ have "{t0--s} \ existence_ivl t0 x0" by (simp add: s segment_subset_existence_ivl iv_defined) have "s \ existence_ivl s (flow t0 x0 s)" by (rule ll_on_open_it.existence_ivl_initial_time; fact) have "{s--t + s} \ existence_ivl s (flow t0 x0 s)" by (rule ll_on_open_it.segment_subset_existence_ivl; fact) have unique: "flow t0 x0 u = flow s (flow t0 x0 s) u" if "u \ {s--t + s}" "u \ {t0--s}" for u using ll_on_open_it_axioms ll_on_open_it.flow_solves_ode[OF ll_on_open_it_axioms iv_defined] ll_on_open_it.flow_solves_ode[OF ll_on_open_it_axioms iv_defined'] s apply (rule ll_on_open_it.unique_on_intersection) using \s \ existence_ivl s (flow t0 x0 s)\ existence_ivl_subset \flow t0 x0 s \ X\ \s \ T\ iv_defined s t ll_on_open_it.in_existence_between_zeroI that ll_on_open_it_axioms ll_on_open_it.mem_existence_ivl_subset by (auto simp: is_interval_existence_ivl) let ?un = "{t0 -- s} \ {s -- t + s}" let ?if = "\t. if t \ {t0 -- s} then flow t0 x0 t else flow s (flow t0 x0 s) t" have "(?if solves_ode (\t. if t \ {t0 -- s} then f t else f t)) ?un (X \ X)" apply (rule connection_solves_ode) subgoal by (rule solves_ode_on_subset[OF flow_solves_ode[OF iv_defined] \{t0--s} \ _\ order_refl]) subgoal by (rule solves_ode_on_subset[OF ll_on_open_it.flow_solves_ode[OF ll_on_open_it_axioms iv_defined'] \{s--t + s} \ _\ order_refl]) subgoal by simp subgoal by simp subgoal by (rule unique) auto subgoal by simp done then have ifsol: "(?if solves_ode f) ?un X" by simp moreover have "?un \ existence_ivl t0 x0" using existence_ivl_subset[of x0] ll_on_open_it.existence_ivl_subset[OF ll_on_open_it_axioms, of s "flow t0 x0 s"] \{t0 -- s} \ _\ \{s--t + s} \ _\ by (intro existence_ivl_maximal_interval[OF ifsol]) (auto intro!: is_real_interval_union) then show "s + t \ existence_ivl t0 x0" by (auto simp: ac_simps) have "(flow t0 x0 solves_ode f) ?un X" using \{t0--s} \ _\ \{s -- t + s} \ _\ by (intro solves_ode_on_subset[OF flow_solves_ode \?un \ _\ order_refl] iv_defined) moreover have "s \ ?un" by simp ultimately have "?if (s + t) = flow t0 x0 (s + t)" apply (rule ll_on_open_it.unique_solution) using existence_ivl_subset[of x0] ll_on_open_it.existence_ivl_subset[OF ll_on_open_it_axioms, of s "flow t0 x0 s"] \{t0 -- s} \ _\ \{s--t + s} \ _\ by (auto intro!: is_real_interval_union simp: ac_simps) with unique[of "s + t"] show "flow t0 x0 (s + t) = flow s (flow t0 x0 s) (s + t)" by (auto split: if_splits simp: ac_simps) qed lemma assumes t: "t \ existence_ivl t0 x0" shows flows_reverse: "flow t (flow t0 x0 t) t0 = x0" and existence_ivl_reverse: "t0 \ existence_ivl t (flow t0 x0 t)" proof - have iv_defined: "t0 \ T" "x0 \ X" using mem_existence_ivl_iv_defined t by blast+ show "t0 \ existence_ivl t (flow t0 x0 t)" using assms by (metis (no_types, opaque_lifting) closed_segment_commute closed_segment_subset_interval ends_in_segment(2) general.csol(2-4) general.existence_ivl_maximal_segment general.is_interval_existence_ivl is_interval_closed_segment_1 iv_defined ll_on_open_it.equals_flowI local.existence_ivl_initial_time local.flow_initial_time local.ll_on_open_it_axioms) then have "flow t (flow t0 x0 t) (t + (t0 - t)) = flow t0 x0 (t + (t0 - t))" by (intro flow_trans[symmetric]) (auto simp: t iv_defined) then show "flow t (flow t0 x0 t) t0 = x0" by (simp add: iv_defined) qed lemma flow_has_derivative: assumes "t \ existence_ivl t0 x0" shows "(flow t0 x0 has_derivative (\i. i *\<^sub>R f t (flow t0 x0 t))) (at t)" proof - have "(flow t0 x0 has_derivative (\i. i *\<^sub>R f t (flow t0 x0 t))) (at t within existence_ivl t0 x0)" using flow_has_vderiv_on by (auto simp: has_vderiv_on_def has_vector_derivative_def assms mem_existence_ivl_iv_defined[OF assms]) then show ?thesis by (simp add: at_within_open[OF assms open_existence_ivl]) qed lemma flow_has_vector_derivative: assumes "t \ existence_ivl t0 x0" shows "(flow t0 x0 has_vector_derivative f t (flow t0 x0 t)) (at t)" using flow_has_derivative[OF assms] by (simp add: has_vector_derivative_def) lemma flow_has_vector_derivative_at_0: assumes"t \ existence_ivl t0 x0" shows "((\h. flow t0 x0 (t + h)) has_vector_derivative f t (flow t0 x0 t)) (at 0)" proof - from flow_has_vector_derivative[OF assms] have "((+) t has_vector_derivative 1) (at 0)" "(flow t0 x0 has_vector_derivative f t (flow t0 x0 t)) (at (t + 0))" by (auto intro!: derivative_eq_intros) from vector_diff_chain_at[OF this] show ?thesis by (simp add: o_def) qed lemma assumes "t \ existence_ivl t0 x0" shows closed_segment_subset_existence_ivl: "closed_segment t0 t \ existence_ivl t0 x0" and ivl_subset_existence_ivl: "{t0 .. t} \ existence_ivl t0 x0" and ivl_subset_existence_ivl': "{t .. t0} \ existence_ivl t0 x0" using assms in_existence_between_zeroI by (auto simp: closed_segment_eq_real_ivl) lemma flow_fixed_point: assumes t: "t \ existence_ivl t0 x0" shows "flow t0 x0 t = x0 + ivl_integral t0 t (\t. f t (flow t0 x0 t))" proof - have "(flow t0 x0 has_vderiv_on (\s. f s (flow t0 x0 s))) {t0 -- t}" using closed_segment_subset_existence_ivl[OF t] by (auto intro!: has_vector_derivative_at_within flow_has_vector_derivative simp: has_vderiv_on_def) from fundamental_theorem_of_calculus_ivl_integral[OF this] have "((\t. f t (flow t0 x0 t)) has_ivl_integral flow t0 x0 t - x0) t0 t" by (simp add: mem_existence_ivl_iv_defined[OF assms]) from this[THEN ivl_integral_unique] show ?thesis by simp qed lemma flow_continuous: "t \ existence_ivl t0 x0 \ continuous (at t) (flow t0 x0)" by (metis has_derivative_continuous flow_has_derivative) lemma flow_tendsto: "t \ existence_ivl t0 x0 \ (ts \ t) F \ ((\s. flow t0 x0 (ts s)) \ flow t0 x0 t) F" by (rule isCont_tendsto_compose[OF flow_continuous]) lemma flow_continuous_on: "continuous_on (existence_ivl t0 x0) (flow t0 x0)" by (auto intro!: flow_continuous continuous_at_imp_continuous_on) lemma flow_continuous_on_intro: "continuous_on s g \ (\xa. xa \ s \ g xa \ existence_ivl t0 x0) \ continuous_on s (\xa. flow t0 x0 (g xa))" by (auto intro!: continuous_on_compose2[OF flow_continuous_on]) lemma f_flow_continuous: assumes "t \ existence_ivl t0 x0" shows "isCont (\t. f t (flow t0 x0 t)) t" by (rule continuous_on_interior) (insert existence_ivl_subset assms, auto intro!: flow_in_domain flow_continuous_on continuous_intros simp: interior_open open_existence_ivl) lemma exponential_initial_condition: assumes y0: "t \ existence_ivl t0 y0" assumes z0: "t \ existence_ivl t0 z0" assumes "Y \ X" assumes remain: "\s. s \ closed_segment t0 t \ flow t0 y0 s \ Y" "\s. s \ closed_segment t0 t \ flow t0 z0 s \ Y" assumes lipschitz: "\s. s \ closed_segment t0 t \ K-lipschitz_on Y (f s)" shows "norm (flow t0 y0 t - flow t0 z0 t) \ norm (y0 - z0) * exp ((K + 1) * abs (t - t0))" proof cases assume "y0 = z0" thus ?thesis by simp next assume ne: "y0 \ z0" define K' where "K' \ K + 1" from lipschitz have "K'-lipschitz_on Y (f s)" if "s \ {t0 -- t}" for s using that by (auto simp: lipschitz_on_def K'_def intro!: order_trans[OF _ mult_right_mono[of K "K + 1"]]) from mem_existence_ivl_iv_defined[OF y0] mem_existence_ivl_iv_defined[OF z0] have "t0 \ T" and inX: "y0 \ X" "z0 \ X" by auto from remain[of t0] inX \t0 \ T \ have "y0 \ Y" "z0 \ Y" by auto define v where "v \ \t. norm (flow t0 y0 t - flow t0 z0 t)" { fix s assume s: "s \ {t0 -- t}" with s closed_segment_subset_existence_ivl[OF y0] closed_segment_subset_existence_ivl[OF z0] have y0': "s \ existence_ivl t0 y0" and z0': "s \ existence_ivl t0 z0" by (auto simp: closed_segment_eq_real_ivl) have integrable: "(\t. f t (flow t0 y0 t)) integrable_on {t0--s}" "(\t. f t (flow t0 z0 t)) integrable_on {t0--s}" using closed_segment_subset_existence_ivl[OF y0'] closed_segment_subset_existence_ivl[OF z0'] \y0 \ X\ \z0 \ X\ \t0 \ T\ by (auto intro!: continuous_at_imp_continuous_on f_flow_continuous integrable_continuous_closed_segment) hence int: "flow t0 y0 s - flow t0 z0 s = y0 - z0 + ivl_integral t0 s (\t. f t (flow t0 y0 t) - f t (flow t0 z0 t))" unfolding v_def using flow_fixed_point[OF y0'] flow_fixed_point[OF z0'] s by (auto simp: algebra_simps ivl_integral_diff) have "v s \ v t0 + K' * integral {t0 -- s} (\t. v t)" using closed_segment_subset_existence_ivl[OF y0'] closed_segment_subset_existence_ivl[OF z0'] s using closed_segment_closed_segment_subset[OF _ _ s, of _ t0, simplified] by (subst integral_mult) (auto simp: integral_mult v_def int inX \t0 \ T\ simp del: Henstock_Kurzweil_Integration.integral_mult_right intro!: norm_triangle_le ivl_integral_norm_bound_integral integrable_continuous_closed_segment continuous_intros continuous_at_imp_continuous_on flow_continuous f_flow_continuous lipschitz_on_normD[OF \_ \ K'-lipschitz_on _ _\] remain) } note le = this have cont: "continuous_on {t0 -- t} v" using closed_segment_subset_existence_ivl[OF y0] closed_segment_subset_existence_ivl[OF z0] inX by (auto simp: v_def \t0 \ T\ intro!: continuous_at_imp_continuous_on continuous_intros flow_continuous) have nonneg: "\t. v t \ 0" by (auto simp: v_def) from ne have pos: "v t0 > 0" by (auto simp: v_def \t0 \ T\ inX) have lippos: "K' > 0" proof - have "0 \ dist (f t0 y0) (f t0 z0)" by simp also from lipschitz_onD[OF lipschitz \y0 \ Y\ \z0 \ Y\, of t0]ne have "\ \ K * dist y0 z0" by simp finally have "0 \ K" by (metis dist_le_zero_iff ne zero_le_mult_iff) thus ?thesis by (simp add: K'_def) qed from le cont nonneg pos \0 < K'\ have "v t \ v t0 * exp (K' * abs (t - t0))" by (rule gronwall_general_segment) simp_all thus ?thesis by (simp add: v_def K'_def \t0 \ T\ inX) qed lemma existence_ivl_cballs: assumes iv_defined: "t0 \ T" "x0 \ X" obtains t u L where "\y. y \ cball x0 u \ cball t0 t \ existence_ivl t0 y" "\s y. y \ cball x0 u \ s \ cball t0 t \ flow t0 y s \ cball y u" "L-lipschitz_on (cball t0 t\cball x0 u) (\(t, x). flow t0 x t)" "\y. y \ cball x0 u \ cball y u \ X" "0 < t" "0 < u" proof - note iv_defined from local_unique_solutions[OF this] obtain t u L where tu: "0 < t" "0 < u" and subsT: "cball t0 t \ existence_ivl t0 x0" and subs': "cball x0 (2 * u) \ X" and lipschitz: "\s. s \ cball t0 t \ L-lipschitz_on (cball x0 (2*u)) (f s)" and usol: "\y. y \ cball x0 u \ (flow t0 y usolves_ode f from t0) (cball t0 t) (cball y u)" and subs: "\y. y \ cball x0 u \ cball y u \ X" by metis { fix y assume y: "y \ cball x0 u" from subs[OF y] \0 < u\ have "y \ X" by auto note iv' = \t0 \ T\ \y \ X\ from usol[OF y, THEN usolves_odeD(1)] have sol1: "(flow t0 y solves_ode f) (cball t0 t) (cball y u)" . from sol1 order_refl subs[OF y] have sol: "(flow t0 y solves_ode f) (cball t0 t) X" by (rule solves_ode_on_subset) note * = maximal_existence_flow[OF sol flow_initial_time is_interval_cball_1 _ order_trans[OF subsT existence_ivl_subset], unfolded centre_in_cball, OF iv' less_imp_le[OF \0 < t\]] have eivl: "cball t0 t \ existence_ivl t0 y" by (rule *) have "flow t0 y s \ cball y u" if "s \ cball t0 t" for s by (rule solves_odeD(2)[OF sol1 that]) note eivl this } note * = this note * moreover have cont_on_f_flow: "\x1 S. S \ cball t0 t \ x1 \ cball x0 u \ continuous_on S (\t. f t (flow t0 x1 t))" using subs[of x0] \u > 0\ *(1) iv_defined by (auto intro!: continuous_at_imp_continuous_on f_flow_continuous) have "bounded ((\(t, x). f t x) ` (cball t0 t \ cball x0 (2 * u)))" using subs' subsT existence_ivl_subset[of x0] by (auto intro!: compact_imp_bounded compact_continuous_image compact_Times continuous_intros simp: split_beta') then obtain B where B: "\s y. s \ cball t0 t \ y \ cball x0 (2 * u) \ norm (f s y) \ B" "B > 0" by (auto simp: bounded_pos cball_def) have flow_in_cball: "flow t0 x1 s \ cball x0 (2 * u)" if s: "s \ cball t0 t" and x1: "x1 \ cball x0 u" for s::real and x1 proof - from *(2)[OF x1 s] have "flow t0 x1 s \ cball x1 u" . also have "\ \ cball x0 (2 * u)" using x1 by (auto intro!: dist_triangle_le[OF add_mono, of _ x1 u _ u, simplified] simp: dist_commute) finally show ?thesis . qed have "(B + exp ((L + 1) * \t\))-lipschitz_on (cball t0 t\cball x0 u) (\(t, x). flow t0 x t)" proof (rule lipschitz_onI, safe) fix t1 t2 :: real and x1 x2 assume t1: "t1 \ cball t0 t" and t2: "t2 \ cball t0 t" and x1: "x1 \ cball x0 u" and x2: "x2 \ cball x0 u" have t1_ex: "t1 \ existence_ivl t0 x1" and t2_ex: "t2 \ existence_ivl t0 x1" "t2 \ existence_ivl t0 x2" and "x1 \ cball x0 (2*u)" "x2 \ cball x0 (2*u)" using *(1)[OF x1] *(1)[OF x2] t1 t2 x1 x2 tu by auto have "dist (flow t0 x1 t1) (flow t0 x2 t2) \ dist (flow t0 x1 t1) (flow t0 x1 t2) + dist (flow t0 x1 t2) (flow t0 x2 t2)" by (rule dist_triangle) also have "dist (flow t0 x1 t2) (flow t0 x2 t2) \ dist x1 x2 * exp ((L + 1) * \t2 - t0\)" unfolding dist_norm proof (rule exponential_initial_condition[where Y = "cball x0 (2 * u)"]) fix s assume "s \ closed_segment t0 t2" hence s: "s \ cball t0 t" using t2 by (auto simp: dist_real_def closed_segment_eq_real_ivl split: if_split_asm) show "flow t0 x1 s \ cball x0 (2 * u)" by (rule flow_in_cball[OF s x1]) show "flow t0 x2 s \ cball x0 (2 * u)" by (rule flow_in_cball[OF s x2]) show "L-lipschitz_on (cball x0 (2 * u)) (f s)" if "s \ closed_segment t0 t2" for s using that centre_in_cball convex_contains_segment less_imp_le t2 tu(1) by (blast intro!: lipschitz) qed (fact)+ also have "\ \ dist x1 x2 * exp ((L + 1) * \t\)" using \u > 0\ t2 by (auto intro!: mult_left_mono add_nonneg_nonneg lipschitz[THEN lipschitz_on_nonneg] simp: cball_eq_empty cball_eq_sing' dist_real_def) also have "x1 \ X" using x1 subs[of x0] \u > 0\ by auto have *: "\t0 - t1\ \ t \ x \ {t0--t1} \ \t0 - x\ \ t" "\t0 - t2\ \ t \ x \ {t0--t2} \ \t0 - x\ \ t" "\t0 - t1\ \ t \ \t0 - t2\ \ t \ x \ {t1--t2} \ \t0 - x\ \ t" for x using t1 t2 t1_ex x1 flow_in_cball[OF _ x1] by (auto simp: closed_segment_eq_real_ivl split: if_splits) have integrable: "(\t. f t (flow t0 x1 t)) integrable_on {t0--t1}" "(\t. f t (flow t0 x1 t)) integrable_on {t0--t2}" "(\t. f t (flow t0 x1 t)) integrable_on {t1--t2}" using t1 t2 t1_ex x1 flow_in_cball[OF _ x1] by (auto intro!: order_trans[OF integral_bound[where B=B]] cont_on_f_flow B integrable_continuous_closed_segment intro: * simp: dist_real_def integral_minus_sets') have *: "\t0 - t1\ \ t \ \t0 - t2\ \ t \ s \ {t1--t2} \ \t0 - s\ \ t" for s by (auto simp: closed_segment_eq_real_ivl split: if_splits) note [simp] = t1_ex t2_ex \x1 \ X\ integrable have "dist (flow t0 x1 t1) (flow t0 x1 t2) \ dist t1 t2 * B" using t1 t2 x1 flow_in_cball[OF _ x1] \t0 \ T\ ivl_integral_combine[of "\t. f t (flow t0 x1 t)" t2 t0 t1] ivl_integral_combine[of "\t. f t (flow t0 x1 t)" t1 t0 t2] by (auto simp: flow_fixed_point dist_norm add.commute closed_segment_commute norm_minus_commute ivl_integral_minus_sets' ivl_integral_minus_sets intro!: order_trans[OF ivl_integral_bound[where B=B]] cont_on_f_flow B dest: *) finally have "dist (flow t0 x1 t1) (flow t0 x2 t2) \ dist t1 t2 * B + dist x1 x2 * exp ((L + 1) * \t\)" by arith also have "\ \ dist (t1, x1) (t2, x2) * B + dist (t1, x1) (t2, x2) * exp ((L + 1) * \t\)" using \B > 0\ by (auto intro!: add_mono mult_right_mono simp: dist_prod_def) finally show "dist (flow t0 x1 t1) (flow t0 x2 t2) \ (B + exp ((L + 1) * \t\)) * dist (t1, x1) (t2, x2)" by (simp add: algebra_simps) qed (simp add: \0 < B\ less_imp_le) ultimately show thesis using subs tu .. qed context fixes x0 assumes iv_defined: "t0 \ T" "x0 \ X" begin lemma existence_ivl_notempty: "existence_ivl t0 x0 \ {}" using existence_ivl_initial_time iv_defined by auto lemma initial_time_bounds: shows "bdd_above (existence_ivl t0 x0) \ t0 < Sup (existence_ivl t0 x0)" (is "?a \ _") and "bdd_below (existence_ivl t0 x0) \ Inf (existence_ivl t0 x0) < t0" (is "?b \ _") proof - from local_unique_solutions[OF iv_defined] obtain te where te: "te > 0" "cball t0 te \ existence_ivl t0 x0" by metis then show "t0 < Sup (existence_ivl t0 x0)" if bdd: "bdd_above (existence_ivl t0 x0)" using less_cSup_iff[OF existence_ivl_notempty bdd, of t0] iv_defined by (auto simp: dist_real_def intro!: bexI[where x="t0 + te"]) from te show "Inf (existence_ivl t0 x0) < t0" if bdd: "bdd_below (existence_ivl t0 x0)" unfolding cInf_less_iff[OF existence_ivl_notempty bdd, of t0] by (auto simp: dist_real_def iv_defined intro!: bexI[where x="t0 - te"]) qed lemma flow_leaves_compact_ivl_right: assumes bdd: "bdd_above (existence_ivl t0 x0)" defines "b \ Sup (existence_ivl t0 x0)" assumes "b \ T" assumes "compact K" assumes "K \ X" obtains t where "t \ t0" "t \ existence_ivl t0 x0" "flow t0 x0 t \ K" proof (atomize_elim, rule ccontr, auto) note iv_defined note ne = existence_ivl_notempty assume K[rule_format]: "\t. t \ existence_ivl t0 x0 \ t0 \ t \ flow t0 x0 t \ K" have b_upper: "t \ b" if "t \ existence_ivl t0 x0" for t unfolding b_def by (rule cSup_upper[OF that bdd]) have less_b_iff: "y < b \ (\x\existence_ivl t0 x0. y < x)" for y unfolding b_def less_cSup_iff[OF ne bdd] .. have "t0 \ b" by (simp add: iv_defined b_upper) then have geI: "t \ {t0-- t0 \ t" for t by (auto simp: half_open_segment_real) have subset: "{t0 --< b} \ existence_ivl t0 x0" using \t0 \ b\ in_existence_between_zeroI by (auto simp: half_open_segment_real iv_defined less_b_iff) have sol: "(flow t0 x0 solves_ode f) {t0 --< b} K" apply (rule solves_odeI) apply (rule has_vderiv_on_subset[OF solves_odeD(1)[OF flow_solves_ode] subset]) using subset iv_defined by (auto intro!: K geI) have cont: "continuous_on ({t0--b} \ K) (\(t, x). f t x)" using \K \ X\ closed_segment_subset_domainI[OF iv_defined(1) \b \ T\] by (auto simp: split_beta intro!: continuous_intros) from initial_time_bounds(1)[OF bdd] have "t0 \ b" by (simp add: b_def) from solves_ode_half_open_segment_continuation[OF sol cont \compact K\ \t0 \ b\] obtain l where lim: "(flow t0 x0 \ l) (at b within {t0--t. if t = b then l else flow t0 x0 t) solves_ode f) {t0--b} K" . have "b \ existence_ivl t0 x0" using \t0 \ b\ closed_segment_subset_domainI[OF \t0 \ T\ \b \ T\] by (intro existence_ivl_maximal_segment[OF solves_ode_on_subset[OF limsol order_refl \K \ X\]]) (auto simp: iv_defined) have "flow t0 x0 b \ X" by (simp add: \b \ existence_ivl t0 x0\ flow_in_domain iv_defined) from ll_on_open_it.local_unique_solutions[OF ll_on_open_it_axioms \b \ T\ \flow t0 x0 b \ X\] obtain e where "e > 0" "cball b e \ existence_ivl b (flow t0 x0 b)" by metis then have "e + b \ existence_ivl b (flow t0 x0 b)" by (auto simp: dist_real_def) from existence_ivl_trans[OF \b \ existence_ivl t0 x0\ \e + b \ existence_ivl _ _\] have "b + e \ existence_ivl t0 x0" . from b_upper[OF this] \e > 0\ show False by simp qed lemma flow_leaves_compact_ivl_left: assumes bdd: "bdd_below (existence_ivl t0 x0)" defines "b \ Inf (existence_ivl t0 x0)" assumes "b \ T" assumes "compact K" assumes "K \ X" obtains t where "t \ t0" "t \ existence_ivl t0 x0" "flow t0 x0 t \ K" proof - interpret rev: ll_on_open "(preflect t0 ` T)" "(\t. - f (preflect t0 t))" X .. from antimono_preflect bdd have bdd_rev: "bdd_above (rev.existence_ivl t0 x0)" unfolding rev_existence_ivl_eq by (rule bdd_above_image_antimono) note ne = existence_ivl_notempty have "Sup (rev.existence_ivl t0 x0) = preflect t0 b" using continuous_at_Inf_antimono[OF antimono_preflect _ ne bdd] by (simp add: continuous_preflect b_def rev_existence_ivl_eq) then have Sup_mem: "Sup (rev.existence_ivl t0 x0) \ preflect t0 ` T" using \b \ T\ by auto have rev_iv: "t0 \ preflect t0 ` T" "x0 \ X" using iv_defined by auto from rev.flow_leaves_compact_ivl_right[OF rev_iv bdd_rev Sup_mem \compact K\ \K \ X\] obtain t where "t0 \ t" "t \ rev.existence_ivl t0 x0" "rev.flow t0 x0 t \ K" . then have "preflect t0 t \ t0" "preflect t0 t \ existence_ivl t0 x0" "flow t0 x0 (preflect t0 t) \ K" by (auto simp: rev_existence_ivl_eq rev_flow_eq) thus ?thesis .. qed lemma sup_existence_maximal: assumes "\t. t0 \ t \ t \ existence_ivl t0 x0 \ flow t0 x0 t \ K" assumes "compact K" "K \ X" assumes "bdd_above (existence_ivl t0 x0)" shows "Sup (existence_ivl t0 x0) \ T" using flow_leaves_compact_ivl_right[of K] assms by force lemma inf_existence_minimal: assumes "\t. t \ t0 \ t \ existence_ivl t0 x0 \ flow t0 x0 t \ K" assumes "compact K" "K \ X" assumes "bdd_below (existence_ivl t0 x0)" shows "Inf (existence_ivl t0 x0) \ T" using flow_leaves_compact_ivl_left[of K] assms by force end lemma subset_mem_compact_implies_subset_existence_interval: assumes ivl: "t0 \ T'" "is_interval T'" "T' \ T" assumes iv_defined: "x0 \ X" assumes mem_compact: "\t. t \ T' \ t \ existence_ivl t0 x0 \ flow t0 x0 t \ K" assumes K: "compact K" "K \ X" shows "T' \ existence_ivl t0 x0" proof (rule ccontr) assume "\ T' \ existence_ivl t0 x0" then obtain t' where t': "t' \ existence_ivl t0 x0" "t' \ T'" by auto from assms have iv_defined: "t0 \ T" "x0 \ X" by auto show False proof (cases rule: not_in_connected_cases[OF connected_existence_ivl t'(1) existence_ivl_notempty[OF iv_defined]]) assume bdd: "bdd_below (existence_ivl t0 x0)" assume t'_lower: "t' \ y" if "y \ existence_ivl t0 x0" for y have i: "Inf (existence_ivl t0 x0) \ T'" using initial_time_bounds[OF iv_defined] iv_defined apply - by (rule mem_is_intervalI[of _ t' t0]) (auto simp: ivl t' bdd intro!: t'_lower cInf_greatest[OF existence_ivl_notempty[OF iv_defined]]) have *: "t \ T'" if "t \ t0" "t \ existence_ivl t0 x0" for t by (rule mem_is_intervalI[OF \is_interval T'\ i \t0 \ T'\]) (auto intro!: cInf_lower that bdd) from inf_existence_minimal[OF iv_defined mem_compact K bdd, OF *] show False using i ivl by auto next assume bdd: "bdd_above (existence_ivl t0 x0)" assume t'_upper: "y \ t'" if "y \ existence_ivl t0 x0" for y have s: "Sup (existence_ivl t0 x0) \ T'" using initial_time_bounds[OF iv_defined] apply - apply (rule mem_is_intervalI[of _ t0 t']) by (auto simp: ivl t' bdd intro!: t'_upper cSup_least[OF existence_ivl_notempty[OF iv_defined]]) have *: "t \ T'" if "t0 \ t" "t \ existence_ivl t0 x0" for t by (rule mem_is_intervalI[OF \is_interval T'\ \t0 \ T'\ s]) (auto intro!: cSup_upper that bdd) from sup_existence_maximal[OF iv_defined mem_compact K bdd, OF *] show False using s ivl by auto qed qed lemma mem_compact_implies_subset_existence_interval: assumes iv_defined: "t0 \ T" "x0 \ X" assumes mem_compact: "\t. t \ T \ t \ existence_ivl t0 x0 \ flow t0 x0 t \ K" assumes K: "compact K" "K \ X" shows "T \ existence_ivl t0 x0" by (rule subset_mem_compact_implies_subset_existence_interval; (fact | rule order_refl interval iv_defined)) lemma global_right_existence_ivl_explicit: assumes "b \ t0" assumes b: "b \ existence_ivl t0 x0" obtains d K where "d > 0" "K > 0" "ball x0 d \ X" "\y. y \ ball x0 d \ b \ existence_ivl t0 y" "\t y. y \ ball x0 d \ t \ {t0 .. b} \ dist (flow t0 x0 t) (flow t0 y t) \ dist x0 y * exp (K * abs (t - t0))" proof - note iv_defined = mem_existence_ivl_iv_defined[OF b] define seg where "seg \ (\t. flow t0 x0 t) ` (closed_segment t0 b)" have [simp]: "x0 \ seg" by (auto simp: seg_def intro!: image_eqI[where x=t0] simp: closed_segment_eq_real_ivl iv_defined) have "seg \ {}" by (auto simp: seg_def closed_segment_eq_real_ivl) moreover have "compact seg" using iv_defined b by (auto simp: seg_def closed_segment_eq_real_ivl intro!: compact_continuous_image continuous_at_imp_continuous_on flow_continuous; metis (erased, opaque_lifting) atLeastAtMost_iff closed_segment_eq_real_ivl closed_segment_subset_existence_ivl contra_subsetD order.trans) moreover note open_domain(2) moreover have "seg \ X" using closed_segment_subset_existence_ivl b by (auto simp: seg_def intro!: flow_in_domain iv_defined) ultimately obtain e where e: "0 < e" "{x. infdist x seg \ e} \ X" thm compact_in_open_separated by (rule compact_in_open_separated) define A where "A \ {x. infdist x seg \ e}" have "A \ X" using e by (simp add: A_def) have mem_existence_ivlI: "\s. t0 \ s \ s \ b \ s \ existence_ivl t0 x0" by (rule in_existence_between_zeroI[OF b]) (auto simp: closed_segment_eq_real_ivl) have "compact A" unfolding A_def by (rule compact_infdist_le) fact+ have "compact {t0 .. b}" "{t0 .. b} \ T" subgoal by simp subgoal using mem_existence_ivlI mem_existence_ivl_subset[of _ x0] iv_defined b ivl_subset_existence_ivl by blast done from lipschitz_on_compact[OF this \compact A\ \A \ X\] obtain K' where K': "\t. t \ {t0 .. b} \ K'-lipschitz_on A (f t)" by metis define K where "K \ K' + 1" have "0 < K" "0 \ K" using assms lipschitz_on_nonneg[OF K', of t0] by (auto simp: K_def) have K: "\t. t \ {t0 .. b} \ K-lipschitz_on A (f t)" unfolding K_def using \_ \ lipschitz_on K' A _\ by (rule lipschitz_on_mono) auto have [simp]: "x0 \ A" using \0 < e\ by (auto simp: A_def) define d where "d \ min e (e * exp (-K * (b - t0)))" hence d: "0 < d" "d \ e" "d \ e * exp (-K * (b - t0))" using e by auto have d_times_exp_le: "d * exp (K * (t - t0)) \ e" if "t0 \ t" "t \ b" for t proof - from that have "d * exp (K * (t - t0)) \ d * exp (K * (b - t0))" using \0 \ K\ \0 < d\ by (auto intro!: mult_left_mono) also have "d * exp (K * (b - t0)) \ e" using d by (auto simp: exp_minus divide_simps) finally show ?thesis . qed have "ball x0 d \ X" using d \A \ X\ by (auto simp: A_def dist_commute intro!: infdist_le2[where a=x0]) note iv_defined { fix y assume y: "y \ ball x0 d" hence "y \ A" using d by (auto simp: A_def dist_commute intro!: infdist_le2[where a=x0]) hence "y \ X" using \A \ X\ by auto note y_iv = \t0 \ T\ \y \ X\ have in_A: "flow t0 y t \ A" if t: "t0 \ t" "t \ existence_ivl t0 y" "t \ b" for t proof (rule ccontr) assume flow_out: "flow t0 y t \ A" obtain t' where t': "t0 \ t'" "t' \ t" "\t. t \ {t0 .. t'} \ flow t0 x0 t \ A" "infdist (flow t0 y t') seg \ e" "\t. t \ {t0 .. t'} \ flow t0 y t \ A" proof - let ?out = "((\t. infdist (flow t0 y t) seg) -` {e..}) \ {t0..t}" have "compact ?out" unfolding compact_eq_bounded_closed proof safe show "bounded ?out" by (auto intro!: bounded_closed_interval) have "continuous_on {t0 .. t} ((\t. infdist (flow t0 y t) seg))" using closed_segment_subset_existence_ivl t iv_defined by (force intro!: continuous_at_imp_continuous_on continuous_intros flow_continuous simp: closed_segment_eq_real_ivl) thus "closed ?out" by (simp add: continuous_on_closed_vimage) qed moreover have "t \ (\t. infdist (flow t0 y t) seg) -` {e..} \ {t0..t}" using flow_out \t0 \ t\ by (auto simp: A_def) hence "?out \ {}" by blast ultimately have "\s\?out. \t\?out. s \ t" by (rule compact_attains_inf) then obtain t' where t': "\s. e \ infdist (flow t0 y s) seg \ t0 \ s \ s \ t \ t' \ s" "e \ infdist (flow t0 y t') seg" "t0 \ t'" "t' \ t" by (auto simp: vimage_def Ball_def) metis have flow_in: "flow t0 x0 s \ A" if s: "s \ {t0 .. t'}" for s proof - from s have "s \ closed_segment t0 b" using \t \ b\ t' by (auto simp: closed_segment_eq_real_ivl) then show ?thesis using s \e > 0\ by (auto simp: seg_def A_def) qed have "flow t0 y t' \ A" if "t' = t0" using y d iv_defined that by (auto simp: A_def \y \ X\ infdist_le2[where a=x0] dist_commute) moreover have "flow t0 y s \ A" if s: "s \ {t0 ..< t'}" for s proof - from s have "s \ closed_segment t0 b" using \t \ b\ t' by (auto simp: closed_segment_eq_real_ivl) from t'(1)[of s] have "t' > s \ t0 \ s \ s \ t \ e > infdist (flow t0 y s) seg" by force then show ?thesis using s t' \e > 0\ by (auto simp: seg_def A_def) qed moreover note left_of_in = this have "closed A" using \compact A\ by (auto simp: compact_eq_bounded_closed) have "((\s. flow t0 y s) \ flow t0 y t') (at_left t')" using closed_segment_subset_existence_ivl[OF t(2)] t' \y \ X\ iv_defined by (intro flow_tendsto) (auto intro!: tendsto_intros simp: closed_segment_eq_real_ivl) with \closed A\ _ _ have "t' \ t0 \ flow t0 y t' \ A" proof (rule Lim_in_closed_set) assume "t' \ t0" hence "t' > t0" using t' by auto hence "eventually (\x. x \ t0) (at_left t')" by (metis eventually_at_left less_imp_le) thus "eventually (\x. flow t0 y x \ A) (at_left t')" unfolding eventually_at_filter by eventually_elim (auto intro!: left_of_in) qed simp ultimately have flow_y_in: "s \ {t0 .. t'} \ flow t0 y s \ A" for s by (cases "s = t'"; fastforce) have "t0 \ t'" "t' \ t" "\t. t \ {t0 .. t'} \ flow t0 x0 t \ A" "infdist (flow t0 y t') seg \ e" "\t. t \ {t0 .. t'} \ flow t0 y t \ A" by (auto intro!: flow_in flow_y_in) fact+ thus ?thesis .. qed { fix s assume s: "s \ {t0 .. t'}" hence "t0 \ s" by simp have "s \ b" using t t' s b by auto hence sx0: "s \ existence_ivl t0 x0" by (simp add: \t0 \ s\ mem_existence_ivlI) have sy: "s \ existence_ivl t0 y" by (meson atLeastAtMost_iff contra_subsetD s t'(1) t'(2) that(2) ivl_subset_existence_ivl) have int: "flow t0 y s - flow t0 x0 s = y - x0 + (integral {t0 .. s} (\t. f t (flow t0 y t)) - integral {t0 .. s} (\t. f t (flow t0 x0 t)))" using iv_defined s unfolding flow_fixed_point[OF sx0] flow_fixed_point[OF sy] by (simp add: algebra_simps ivl_integral_def) have "norm (flow t0 y s - flow t0 x0 s) \ norm (y - x0) + norm (integral {t0 .. s} (\t. f t (flow t0 y t)) - integral {t0 .. s} (\t. f t (flow t0 x0 t)))" unfolding int by (rule norm_triangle_ineq) also have "norm (integral {t0 .. s} (\t. f t (flow t0 y t)) - integral {t0 .. s} (\t. f t (flow t0 x0 t))) = norm (integral {t0 .. s} (\t. f t (flow t0 y t) - f t (flow t0 x0 t)))" using closed_segment_subset_existence_ivl[of s x0] sx0 closed_segment_subset_existence_ivl[of s y] sy by (subst Henstock_Kurzweil_Integration.integral_diff) (auto intro!: integrable_continuous_real continuous_at_imp_continuous_on f_flow_continuous simp: closed_segment_eq_real_ivl) also have "\ \ (integral {t0 .. s} (\t. norm (f t (flow t0 y t) - f t (flow t0 x0 t))))" using closed_segment_subset_existence_ivl[of s x0] sx0 closed_segment_subset_existence_ivl[of s y] sy by (intro integral_norm_bound_integral) (auto intro!: integrable_continuous_real continuous_at_imp_continuous_on f_flow_continuous continuous_intros simp: closed_segment_eq_real_ivl) also have "\ \ (integral {t0 .. s} (\t. K * norm ((flow t0 y t) - (flow t0 x0 t))))" using closed_segment_subset_existence_ivl[of s x0] sx0 closed_segment_subset_existence_ivl[of s y] sy iv_defined s t'(3,5) \s \ b\ by (auto simp del: Henstock_Kurzweil_Integration.integral_mult_right intro!: integral_le integrable_continuous_real continuous_at_imp_continuous_on lipschitz_on_normD[OF K] flow_continuous f_flow_continuous continuous_intros simp: closed_segment_eq_real_ivl) also have "\ = K * integral {t0 .. s} (\t. norm (flow t0 y t - flow t0 x0 t))" using closed_segment_subset_existence_ivl[of s x0] sx0 closed_segment_subset_existence_ivl[of s y] sy by (subst integral_mult) (auto intro!: integrable_continuous_real continuous_at_imp_continuous_on lipschitz_on_normD[OF K] flow_continuous f_flow_continuous continuous_intros simp: closed_segment_eq_real_ivl) finally have norm: "norm (flow t0 y s - flow t0 x0 s) \ norm (y - x0) + K * integral {t0 .. s} (\t. norm (flow t0 y t - flow t0 x0 t))" by arith note norm \s \ b\ sx0 sy } note norm_le = this from norm_le(2) t' have "t' \ closed_segment t0 b" by (auto simp: closed_segment_eq_real_ivl) hence "infdist (flow t0 y t') seg \ dist (flow t0 y t') (flow t0 x0 t')" by (auto simp: seg_def infdist_le) also have "\ \ norm (flow t0 y t' - flow t0 x0 t')" by (simp add: dist_norm) also have "\ \ norm (y - x0) * exp (K * \t' - t0\)" unfolding K_def apply (rule exponential_initial_condition[OF _ _ _ _ _ K']) subgoal by (metis atLeastAtMost_iff local.norm_le(4) order_refl t'(1)) subgoal by (metis atLeastAtMost_iff local.norm_le(3) order_refl t'(1)) subgoal using e by (simp add: A_def) subgoal by (metis closed_segment_eq_real_ivl t'(1,5)) subgoal by (metis closed_segment_eq_real_ivl t'(1,3)) subgoal by (simp add: closed_segment_eq_real_ivl local.norm_le(2) t'(1)) done also have "\ < d * exp (K * (t - t0))" using y d t' t by (intro mult_less_le_imp_less) (auto simp: dist_norm[symmetric] dist_commute intro!: mult_mono \0 \ K\) also have "\ \ e" by (rule d_times_exp_le; fact) finally have "infdist (flow t0 y t') seg < e" . with \infdist (flow t0 y t') seg \ e\ show False by (auto simp: frontier_def) qed have "{t0..b} \ existence_ivl t0 y" by (rule subset_mem_compact_implies_subset_existence_interval[OF _ is_interval_cc \{t0..b} \ T\ \y \ X\ in_A \compact A\ \A \ X\]) (auto simp: \t0 \ b\) with \t0 \ b\ have b_in: "b \ existence_ivl t0 y" by force { fix t assume t: "t \ {t0 .. b}" also have "{t0 .. b} = {t0 -- b}" by (auto simp: closed_segment_eq_real_ivl assms) also note closed_segment_subset_existence_ivl[OF b_in] finally have t_in: "t \ existence_ivl t0 y" . note t also note \{t0 .. b} = {t0 -- b}\ also note closed_segment_subset_existence_ivl[OF assms(2)] finally have t_in': "t \ existence_ivl t0 x0" . have "norm (flow t0 y t - flow t0 x0 t) \ norm (y - x0) * exp (K * \t - t0\)" unfolding K_def using t closed_segment_subset_existence_ivl[OF b_in] \0 < e\ by (intro in_A exponential_initial_condition[OF t_in t_in' \A \ X\ _ _ K']) (auto simp: closed_segment_eq_real_ivl A_def seg_def) hence "dist (flow t0 x0 t) (flow t0 y t) \ dist x0 y * exp (K * \t - t0\)" by (auto simp: dist_norm[symmetric] dist_commute) } note b_in this } from \d > 0\ \K > 0\ \ball x0 d \ X\ this show ?thesis .. qed lemma global_left_existence_ivl_explicit: assumes "b \ t0" assumes b: "b \ existence_ivl t0 x0" assumes iv_defined: "t0 \ T" "x0 \ X" obtains d K where "d > 0" "K > 0" "ball x0 d \ X" "\y. y \ ball x0 d \ b \ existence_ivl t0 y" "\t y. y \ ball x0 d \ t \ {b .. t0} \ dist (flow t0 x0 t) (flow t0 y t) \ dist x0 y * exp (K * abs (t - t0))" proof - interpret rev: ll_on_open "(preflect t0 ` T)" "(\t. - f (preflect t0 t))" X .. have t0': "t0 \ preflect t0 ` T" "x0 \ X" by (auto intro!: iv_defined) from assms have "preflect t0 b \ t0" "preflect t0 b \ rev.existence_ivl t0 x0" by (auto simp: rev_existence_ivl_eq) from rev.global_right_existence_ivl_explicit[OF this] obtain d K where dK: "d > 0" "K > 0" "ball x0 d \ X" "\y. y \ ball x0 d \ preflect t0 b \ rev.existence_ivl t0 y" "\t y. y \ ball x0 d \ t \ {t0 .. preflect t0 b} \ dist (rev.flow t0 x0 t) (rev.flow t0 y t) \ dist x0 y * exp (K * abs (t - t0))" by (auto simp: rev_flow_eq \x0 \ X\) have ex_ivlI: "dist x0 y < d \ t \ existence_ivl t0 y" if "b \ t" "t \ t0" for t y using that dK(4)[of y] dK(3) iv_defined by (auto simp: subset_iff rev_existence_ivl_eq[of ] closed_segment_eq_real_ivl iv_defined in_existence_between_zeroI) have "b \ existence_ivl t0 y" if "dist x0 y < d" for y using that dK by (subst existence_ivl_eq_rev) (auto simp: iv_defined intro!: image_eqI[where x="preflect t0 b"]) with dK have "d > 0" "K > 0" "ball x0 d \ X" "\y. y \ ball x0 d \ b \ existence_ivl t0 y" "\t y. y \ ball x0 d \ t \ {b .. t0} \ dist (flow t0 x0 t) (flow t0 y t) \ dist x0 y * exp (K * abs (t - t0))" by (auto simp: flow_eq_rev iv_defined ex_ivlI \x0 \ X\ subset_iff intro!: order_trans[OF dK(5)] image_eqI[where x="preflect t0 b"]) then show ?thesis .. qed lemma global_existence_ivl_explicit: assumes a: "a \ existence_ivl t0 x0" assumes b: "b \ existence_ivl t0 x0" assumes le: "a \ b" obtains d K where "d > 0" "K > 0" "ball x0 d \ X" "\y. y \ ball x0 d \ a \ existence_ivl t0 y" "\y. y \ ball x0 d \ b \ existence_ivl t0 y" "\t y. y \ ball x0 d \ t \ {a .. b} \ dist (flow t0 x0 t) (flow t0 y t) \ dist x0 y * exp (K * abs (t - t0))" proof - note iv_defined = mem_existence_ivl_iv_defined[OF a] define r where "r \ Max {t0, a, b}" define l where "l \ Min {t0, a, b}" have r: "r \ t0" "r \ existence_ivl t0 x0" using a b by (auto simp: max_def r_def iv_defined) obtain dr Kr where right: "0 < dr" "0 < Kr" "ball x0 dr \ X" "\y. y \ ball x0 dr \ r \ existence_ivl t0 y" "\y t. y \ ball x0 dr \ t \ {t0..r} \ dist (flow t0 x0 t) (flow t0 y t) \ dist x0 y * exp (Kr * \t - t0\)" by (rule global_right_existence_ivl_explicit[OF r]) blast have l: "l \ t0" "l \ existence_ivl t0 x0" using a b by (auto simp: min_def l_def iv_defined) obtain dl Kl where left: "0 < dl" "0 < Kl" "ball x0 dl \ X" "\y. y \ ball x0 dl \ l \ existence_ivl t0 y" "\y t. y \ ball x0 dl \ t \ {l .. t0} \ dist (flow t0 x0 t) (flow t0 y t) \ dist x0 y * exp (Kl * \t - t0\)" by (rule global_left_existence_ivl_explicit[OF l iv_defined]) blast define d where "d \ min dr dl" define K where "K \ max Kr Kl" note iv_defined have "0 < d" "0 < K" "ball x0 d \ X" using left right by (auto simp: d_def K_def) moreover { fix y assume y: "y \ ball x0 d" hence "y \ X" using \ball x0 d \ X\ by auto from y closed_segment_subset_existence_ivl[OF left(4), of y] closed_segment_subset_existence_ivl[OF right(4), of y] have "a \ existence_ivl t0 y" "b \ existence_ivl t0 y" by (auto simp: d_def l_def r_def min_def max_def closed_segment_eq_real_ivl split: if_split_asm) } moreover { fix t y assume y: "y \ ball x0 d" and t: "t \ {a .. b}" from y have "y \ X" using \ball x0 d \ X\ by auto have "dist (flow t0 x0 t) (flow t0 y t) \ dist x0 y * exp (K * abs (t - t0))" proof cases assume "t \ t0" hence "dist (flow t0 x0 t) (flow t0 y t) \ dist x0 y * exp (Kr * abs (t - t0))" using y t by (intro right) (auto simp: d_def r_def) also have "exp (Kr * abs (t - t0)) \ exp (K * abs (t - t0))" by (auto simp: mult_left_mono K_def max_def mult_right_mono) finally show ?thesis by (simp add: mult_left_mono) next assume "\t \ t0" hence "dist (flow t0 x0 t) (flow t0 y t) \ dist x0 y * exp (Kl * abs (t - t0))" using y t by (intro left) (auto simp: d_def l_def) also have "exp (Kl * abs (t - t0)) \ exp (K * abs (t - t0))" by (auto simp: mult_left_mono K_def max_def mult_right_mono) finally show ?thesis by (simp add: mult_left_mono) qed } ultimately show ?thesis .. qed lemma eventually_exponential_separation: assumes a: "a \ existence_ivl t0 x0" assumes b: "b \ existence_ivl t0 x0" assumes le: "a \ b" obtains K where "K > 0" "\\<^sub>F y in at x0. \t\{a..b}. dist (flow t0 x0 t) (flow t0 y t) \ dist x0 y * exp (K * \t - t0\)" proof - from global_existence_ivl_explicit[OF assms] obtain d K where *: "d > 0" "K > 0" "ball x0 d \ X" "\y. y \ ball x0 d \ a \ existence_ivl t0 y" "\y. y \ ball x0 d \ b \ existence_ivl t0 y" "\t y. y \ ball x0 d \ t \ {a .. b} \ dist (flow t0 x0 t) (flow t0 y t) \ dist x0 y * exp (K * abs (t - t0))" by auto note \K > 0\ moreover have "eventually (\y. y \ ball x0 d) (at x0)" using \d > 0\[THEN eventually_at_ball] by eventually_elim simp hence "eventually (\y. \t\{a..b}. dist (flow t0 x0 t) (flow t0 y t) \ dist x0 y * exp (K * \t - t0\)) (at x0)" by eventually_elim (safe intro!: *) ultimately show ?thesis .. qed lemma eventually_mem_existence_ivl: assumes b: "b \ existence_ivl t0 x0" shows "\\<^sub>F x in at x0. b \ existence_ivl t0 x" proof - from mem_existence_ivl_iv_defined[OF b] have iv_defined: "t0 \ T" "x0 \ X" by simp_all note eiit = existence_ivl_initial_time[OF iv_defined] { fix a b assume assms: "a \ existence_ivl t0 x0" "b \ existence_ivl t0 x0" "a \ b" from global_existence_ivl_explicit[OF assms] obtain d K where *: "d > 0" "K > 0" "ball x0 d \ X" "\y. y \ ball x0 d \ a \ existence_ivl t0 y" "\y. y \ ball x0 d \ b \ existence_ivl t0 y" "\t y. y \ ball x0 d \ t \ {a .. b} \ dist (flow t0 x0 t) (flow t0 y t) \ dist x0 y * exp (K * abs (t - t0))" by auto have "eventually (\y. y \ ball x0 d) (at x0)" using \d > 0\[THEN eventually_at_ball] by eventually_elim simp then have "\\<^sub>F x in at x0. a \ existence_ivl t0 x \ b \ existence_ivl t0 x" by (eventually_elim) (auto intro!: *) } from this[OF b eiit] this[OF eiit b] show ?thesis by (cases "t0 \ b") (auto simp: eventually_mono) qed lemma uniform_limit_flow: assumes a: "a \ existence_ivl t0 x0" assumes b: "b \ existence_ivl t0 x0" assumes le: "a \ b" shows "uniform_limit {a .. b} (flow t0) (flow t0 x0) (at x0)" proof (rule uniform_limitI) fix e::real assume "0 < e" from eventually_exponential_separation[OF assms] obtain K where "0 < K" "\\<^sub>F y in at x0. \t\{a..b}. dist (flow t0 x0 t) (flow t0 y t) \ dist x0 y * exp (K * \t - t0\)" by auto note this(2) moreover let ?m = "max (abs (b - t0)) (abs (a - t0))" have "eventually (\y. \t\{a..b}. dist x0 y * exp (K * \t - t0\) \ dist x0 y * exp (K * ?m)) (at x0)" using \a \ b\ \0 < K\ by (auto intro!: mult_left_mono always_eventually) moreover have "eventually (\y. dist x0 y * exp (K * ?m) < e) (at x0)" using \0 < e\ by (auto intro!: order_tendstoD tendsto_eq_intros) ultimately show "eventually (\y. \t\{a..b}. dist (flow t0 y t) (flow t0 x0 t) < e) (at x0)" by eventually_elim (force simp: dist_commute) qed lemma eventually_at_fst: assumes "eventually P (at (fst x))" assumes "P (fst x)" shows "eventually (\h. P (fst h)) (at x)" using assms unfolding eventually_at_topological by (metis open_vimage_fst rangeI range_fst vimageE vimageI) lemma eventually_at_snd: assumes "eventually P (at (snd x))" assumes "P (snd x)" shows "eventually (\h. P (snd h)) (at x)" using assms unfolding eventually_at_topological by (metis open_vimage_snd rangeI range_snd vimageE vimageI) lemma shows open_state_space: "open (Sigma X (existence_ivl t0))" and flow_continuous_on_state_space: "continuous_on (Sigma X (existence_ivl t0)) (\(x, t). flow t0 x t)" proof (safe intro!: topological_space_class.openI continuous_at_imp_continuous_on) fix t x assume "x \ X" and t: "t \ existence_ivl t0 x" have iv_defined: "t0 \ T" "x \ X" using mem_existence_ivl_iv_defined[OF t] by auto from \x \ X\ t open_existence_ivl obtain e where e: "e > 0" "cball t e \ existence_ivl t0 x" by (metis open_contains_cball) hence ivl: "t - e \ existence_ivl t0 x" "t + e \ existence_ivl t0 x" "t - e \ t + e" by (auto simp: cball_def dist_real_def) obtain d K where dK: "0 < d" "0 < K" "ball x d \ X" "\y. y \ ball x d \ t - e \ existence_ivl t0 y" "\y. y \ ball x d \ t + e \ existence_ivl t0 y" "\y s. y \ ball x d \ s \ {t - e..t + e} \ dist (flow t0 x s) (flow t0 y s) \ dist x y * exp (K * \s - t0\)" by (rule global_existence_ivl_explicit[OF ivl]) blast let ?T = "ball x d \ ball t e" have "open ?T" by (auto intro!: open_Times) moreover have "(x, t) \ ?T" by (auto simp: dK \0 < e\) moreover have "?T \ Sigma X (existence_ivl t0)" proof safe fix s y assume y: "y \ ball x d" and s: "s \ ball t e" with \ball x d \ X\ show "y \ X" by auto have "ball t e \ closed_segment t0 (t - e) \ closed_segment t0 (t + e)" by (auto simp: closed_segment_eq_real_ivl dist_real_def) with \y \ X\ s closed_segment_subset_existence_ivl[OF dK(4)[OF y]] closed_segment_subset_existence_ivl[OF dK(5)[OF y]] show "s \ existence_ivl t0 y" by auto qed ultimately show "\T. open T \ (x, t) \ T \ T \ Sigma X (existence_ivl t0)" by blast have **: "\\<^sub>F s in at 0. norm (flow t0 (x + fst s) (t + snd s) - flow t0 x t) < 2 * eps" if "eps > 0" for eps :: real proof - have "\\<^sub>F s in at 0. norm (flow t0 (x + fst s) (t + snd s) - flow t0 x t) = norm (flow t0 (x + fst s) (t + snd s) - flow t0 x (t + snd s) + (flow t0 x (t + snd s) - flow t0 x t))" by auto moreover have "\\<^sub>F s in at 0. norm (flow t0 (x + fst s) (t + snd s) - flow t0 x (t + snd s) + (flow t0 x (t + snd s) - flow t0 x t)) \ norm (flow t0 (x + fst s) (t + snd s) - flow t0 x (t + snd s)) + norm (flow t0 x (t + snd s) - flow t0 x t)" by eventually_elim (rule norm_triangle_ineq) moreover have "\\<^sub>F s in at 0. t + snd s \ ball t e" by (auto simp: dist_real_def intro!: order_tendstoD[OF _ \0 < e\] intro!: tendsto_eq_intros) moreover from uniform_limit_flow[OF ivl, THEN uniform_limitD, OF \eps > 0\] have "\\<^sub>F (h::(_ )) in at (fst (0::'a*real)). \t\{t - e..t + e}. dist (flow t0 x t) (flow t0 (x + h) t) < eps" by (subst (asm) at_to_0) (auto simp: eventually_filtermap dist_commute ac_simps) hence "\\<^sub>F (h::(_ * real)) in at 0. \t\{t - e..t + e}. dist (flow t0 x t) (flow t0 (x + fst h) t) < eps" by (rule eventually_at_fst) (simp add: \eps > 0\) moreover have "\\<^sub>F h in at (snd (0::'a * _)). norm (flow t0 x (t + h) - flow t0 x t) < eps" using flow_continuous[OF t, unfolded isCont_def, THEN tendstoD, OF \eps > 0\] by (subst (asm) at_to_0) (auto simp: eventually_filtermap dist_norm ac_simps) hence "\\<^sub>F h::('a * _) in at 0. norm (flow t0 x (t + snd h) - flow t0 x t) < eps" by (rule eventually_at_snd) (simp add: \eps > 0\) ultimately show ?thesis proof eventually_elim case (elim s) note elim(1) also note elim(2) also note elim(5) also from elim(3) have "t + snd s \ {t - e..t + e}" by (auto simp: dist_real_def algebra_simps) from elim(4)[rule_format, OF this] have "norm (flow t0 (x + fst s) (t + snd s) - flow t0 x (t + snd s)) < eps" by (auto simp: dist_commute dist_norm[symmetric]) finally show ?case by simp qed qed have *: "\\<^sub>F s in at 0. norm (flow t0 (x + fst s) (t + snd s) - flow t0 x t) < eps" if "eps > 0" for eps::real proof - from that have "eps / 2 > 0" by simp from **[OF this] show ?thesis by auto qed show "isCont (\(x, y). flow t0 x y) (x, t)" unfolding isCont_iff by (rule LIM_zero_cancel) (auto simp: split_beta' norm_conv_dist[symmetric] intro!: tendstoI *) qed lemmas flow_continuous_on_compose[continuous_intros] = continuous_on_compose_Pair[OF flow_continuous_on_state_space] lemma flow_isCont_state_space: "t \ existence_ivl t0 x0 \ isCont (\(x, t). flow t0 x t) (x0, t)" using flow_continuous_on_state_space[of] mem_existence_ivl_iv_defined[of t x0] using continuous_on_eq_continuous_at open_state_space by fastforce lemma flow_absolutely_integrable_on[integrable_on_simps]: assumes "s \ existence_ivl t0 x0" shows "(\x. norm (flow t0 x0 x)) integrable_on closed_segment t0 s" using assms by (auto simp: closed_segment_eq_real_ivl intro!: integrable_continuous_real continuous_intros flow_continuous_on_intro intro: in_existence_between_zeroI) lemma existence_ivl_eq_domain: assumes iv_defined: "t0 \ T" "x0 \ X" assumes bnd: "\tm tM t x. tm \ T \ tM \ T \ \M. \L. \t \ {tm .. tM}. \x \ X. norm (f t x) \ M + L * norm x" assumes "is_interval T" "X = UNIV" shows "existence_ivl t0 x0 = T" proof - from assms have XI: "x \ X" for x by auto { fix tm tM assume tm: "tm \ T" and tM: "tM \ T" and tmtM: "tm \ t0" "t0 \ tM" from bnd[OF tm tM] obtain M' L' where bnd': "\x t. x \ X \ tm \ t \ t \ tM \ norm (f t x) \ M' + L' * norm x" by force define M where "M \ norm M' + 1" define L where "L \ norm L' + 1" have bnd: "\x t. x \ X \ tm \ t \ t \ tM \ norm (f t x) \ M + L * norm x" by (auto simp: M_def L_def intro!: bnd'[THEN order_trans] add_mono mult_mono) have "M > 0" "L > 0" by (auto simp: L_def M_def) let ?r = "(norm x0 + \tm - tM\ * M + 1) * exp (L * \tm - tM\)" define K where "K \ cball (0::'a) ?r" have K: "compact K" "K \ X" by (auto simp: K_def \X = UNIV\) { fix t assume t: "t \ existence_ivl t0 x0" and le: "tm \ t" "t \ tM" { fix s assume sc: "s \ closed_segment t0 t" then have s: "s \ existence_ivl t0 x0" and le: "tm \ s" "s \ tM" using t le sc using closed_segment_subset_existence_ivl apply - subgoal by force subgoal by (metis (full_types) atLeastAtMost_iff closed_segment_eq_real_ivl order_trans tmtM(1)) subgoal by (metis (full_types) atLeastAtMost_iff closed_segment_eq_real_ivl order_trans tmtM(2)) done from sc have nle: "norm (t0 - s) \ norm (t0 - t)" by (auto simp: closed_segment_eq_real_ivl split: if_split_asm) from flow_fixed_point[OF s] have "norm (flow t0 x0 s) \ norm x0 + integral (closed_segment t0 s) (\t. M + L * norm (flow t0 x0 t))" using tmtM using closed_segment_subset_existence_ivl[OF s] le by (auto simp: intro!: norm_triangle_le norm_triangle_ineq4[THEN order.trans] ivl_integral_norm_bound_integral bnd integrable_continuous_closed_segment integrable_continuous_real continuous_intros continuous_on_subset[OF flow_continuous_on] flow_in_domain mem_existence_ivl_subset) (auto simp: closed_segment_eq_real_ivl split: if_splits) also have "\ = norm x0 + norm (t0 - s) * M + L * integral (closed_segment t0 s) (\t. norm (flow t0 x0 t))" by (simp add: integral_add integrable_on_simps \s \ existence_ivl _ _\ integral_const_closed_segment abs_minus_commute) also have "norm (t0 - s) * M \ norm (t0 - t) * M " using nle \M > 0\ by auto also have "\ \ \ + 1" by simp finally have "norm (flow t0 x0 s) \ norm x0 + norm (t0 - t) * M + 1 + L * integral (closed_segment t0 s) (\t. norm (flow t0 x0 t))" by simp } then have "norm (flow t0 x0 t) \ (norm x0 + norm (t0 - t) * M + 1) * exp (L * \t - t0\)" using closed_segment_subset_existence_ivl[OF t] by (intro gronwall_more_general_segment[where a=t0 and b = t and t = t]) (auto simp: \0 < L\ \0 < M\ less_imp_le intro!: add_nonneg_pos mult_nonneg_nonneg add_nonneg_nonneg continuous_intros flow_continuous_on_intro) also have "\ \ ?r" using le tmtM by (auto simp: less_imp_le \0 < M\ \0 < L\ abs_minus_commute intro!: mult_mono) finally have "flow t0 x0 t \ K" by (simp add: dist_norm K_def) } note flow_compact = this have "{tm..tM} \ existence_ivl t0 x0" using tmtM tm \x0 \ X\ \compact K\ \K \ X\ mem_is_intervalI[OF \is_interval T\ \tm \ T\ \tM \ T\] by (intro subset_mem_compact_implies_subset_existence_interval[OF _ _ _ _flow_compact]) (auto simp: tmtM is_interval_cc) } note bnds = this show "existence_ivl t0 x0 = T" proof safe fix x assume x: "x \ T" from bnds[OF x iv_defined(1)] bnds[OF iv_defined(1) x] show "x \ existence_ivl t0 x0" by (cases "x \ t0") auto qed (insert existence_ivl_subset, fastforce) qed lemma flow_unique: assumes "t \ existence_ivl t0 x0" assumes "phi t0 = x0" assumes "\t. t \ existence_ivl t0 x0 \ (phi has_vector_derivative f t (phi t)) (at t)" assumes "\t. t \ existence_ivl t0 x0 \ phi t \ X" shows "flow t0 x0 t = phi t" apply (rule maximal_existence_flow[where K="existence_ivl t0 x0"]) subgoal by (auto intro!: solves_odeI simp: has_vderiv_on_def assms at_within_open[OF _ open_existence_ivl]) subgoal by fact subgoal by simp subgoal using mem_existence_ivl_iv_defined[OF \t \ existence_ivl t0 x0\] by simp subgoal by (simp add: existence_ivl_subset) subgoal by fact done lemma flow_unique_on: assumes "t \ existence_ivl t0 x0" assumes "phi t0 = x0" assumes "(phi has_vderiv_on (\t. f t (phi t))) (existence_ivl t0 x0)" assumes "\t. t \ existence_ivl t0 x0 \ phi t \ X" shows "flow t0 x0 t = phi t" using flow_unique[where phi=phi, OF assms(1,2) _ assms(4)] assms(3) by (auto simp: has_vderiv_on_open) end \ \@{thm local_lipschitz}\ locale two_ll_on_open = F: ll_on_open T1 F X + G: ll_on_open T2 G X for F T1 G T2 X J x0 + fixes e::real and K assumes t0_in_J: "0 \ J" assumes J_subset: "J \ F.existence_ivl 0 x0" assumes J_ivl: "is_interval J" assumes F_lipschitz: "\t. t \ J \ K-lipschitz_on X (F t)" assumes K_pos: "0 < K" assumes F_G_norm_ineq: "\t x. t \ J \ x \ X \ norm (F t x - G t x) < e" begin context begin lemma F_iv_defined: "0 \ T1" "x0 \ X" subgoal using F.existence_ivl_initial_time_iff J_subset t0_in_J by blast subgoal using F.mem_existence_ivl_iv_defined(2) J_subset t0_in_J by blast done lemma e_pos: "0 < e" using le_less_trans[OF norm_ge_zero F_G_norm_ineq[OF t0_in_J F_iv_defined(2)]] by assumption qualified definition "flow0 t = F.flow 0 x0 t" qualified definition "Y t = G.flow 0 x0 t" lemma norm_X_Y_bound: shows "\t \ J \ G.existence_ivl 0 x0. norm (flow0 t - Y t) \ e / K * (exp(K * \t\) - 1)" proof(safe) fix t assume "t \ J" assume tG: "t \ G.existence_ivl 0 x0" have "0 \ J" by (simp add: t0_in_J) let ?u="\t. norm (flow0 t - Y t)" show "norm (flow0 t - Y t) \ e / K * (exp (K * \t\) - 1)" proof(cases "0 \ t") assume "0 \ t" hence [simp]: "\t\ = t" by simp have t0_t_in_J: "{0..t} \ J" using \t \ J\ \0 \ J\ J_ivl using mem_is_interval_1_I atLeastAtMost_iff subsetI by blast note F_G_flow_cont[continuous_intros] = continuous_on_subset[OF F.flow_continuous_on] continuous_on_subset[OF G.flow_continuous_on] have "?u t + e/K \ e/K * exp(K * t)" proof(rule gronwall[where g="\t. ?u t + e/K", OF _ _ _ _ K_pos \0 \ t\ order.refl]) fix s assume "0 \ s" "s \ t" hence "{0..s} \ J" using t0_t_in_J by auto hence t0_s_in_existence: "{0..s} \ F.existence_ivl 0 x0" "{0..s} \ G.existence_ivl 0 x0" using J_subset tG \0 \ s\ \s \ t\ G.ivl_subset_existence_ivl[OF tG] by auto hence s_in_existence: "s \ F.existence_ivl 0 x0" "s \ G.existence_ivl 0 x0" using \0 \ s\ by auto note cont_statements[continuous_intros] = F_iv_defined (* G.iv_defined *) F.flow_in_domain G.flow_in_domain F.mem_existence_ivl_subset G.mem_existence_ivl_subset have [integrable_on_simps]: "continuous_on {0..s} (\s. F s (F.flow 0 x0 s))" "continuous_on {0..s} (\s. G s (G.flow 0 x0 s))" "continuous_on {0..s} (\s. F s (G.flow 0 x0 s))" "continuous_on {0..s} (\s. G s (F.flow 0 x0 s))" using t0_s_in_existence by (auto intro!: continuous_intros integrable_continuous_real) have "flow0 s - Y s = integral {0..s} (\s. F s (flow0 s) - G s (Y s))" using \0 \ s\ by (simp add: flow0_def Y_def Henstock_Kurzweil_Integration.integral_diff integrable_on_simps ivl_integral_def F.flow_fixed_point[OF s_in_existence(1)] G.flow_fixed_point[OF s_in_existence(2)]) also have "... = integral {0..s} (\s. (F s (flow0 s) - F s (Y s)) + (F s (Y s) - G s (Y s)))" by simp also have "... = integral {0..s} (\s. F s (flow0 s) - F s (Y s)) + integral {0..s} (\s. F s (Y s) - G s (Y s))" by (simp add: Henstock_Kurzweil_Integration.integral_diff Henstock_Kurzweil_Integration.integral_add flow0_def Y_def integrable_on_simps) finally have "?u s \ norm (integral {0..s} (\s. F s (flow0 s) - F s (Y s))) + norm (integral {0..s} (\s. F s (Y s) - G s (Y s)))" by (simp add: norm_triangle_ineq) also have "... \ integral {0..s} (\s. norm (F s (flow0 s) - F s (Y s))) + integral {0..s} (\s. norm (F s (Y s) - G s (Y s)))" using t0_s_in_existence by (auto simp add: flow0_def Y_def intro!: add_mono continuous_intros continuous_on_imp_absolutely_integrable_on) also have "... \ integral {0..s} (\s. K * ?u s) + integral {0..s} (\s. e)" proof (rule add_mono[OF integral_le integral_le]) show "norm (F x (flow0 x) - F x (Y x)) \ K * norm (flow0 x - Y x)" if "x \ {0..s}" for x using F_lipschitz[unfolded lipschitz_on_def, THEN conjunct2] that cont_statements(1,2,4) t0_s_in_existence F_iv_defined (* G.iv_defined *) by (metis F_lipschitz flow0_def Y_def \{0..s} \ J\ lipschitz_on_normD F.flow_in_domain G.flow_in_domain subsetCE) show "\x. x \ {0..s} \ norm (F x (Y x) - G x (Y x)) \ e" using F_G_norm_ineq cont_statements(2,3) t0_s_in_existence using Y_def \{0..s} \ J\ cont_statements(5) subset_iff G.flow_in_domain by (metis eucl_less_le_not_le) qed (simp_all add: t0_s_in_existence continuous_intros integrable_on_simps flow0_def Y_def) also have "... = K * integral {0..s} (\s. ?u s + e / K)" using K_pos t0_s_in_existence by (simp_all add: algebra_simps Henstock_Kurzweil_Integration.integral_add flow0_def Y_def continuous_intros continuous_on_imp_absolutely_integrable_on) finally show "?u s + e / K \ e / K + K * integral {0..s} (\s. ?u s + e / K)" by simp next show "continuous_on {0..t} (\t. norm (flow0 t - Y t) + e / K)" using t0_t_in_J J_subset G.ivl_subset_existence_ivl[OF tG] by (auto simp add: flow0_def Y_def intro!: continuous_intros) next fix s assume "0 \ s" "s \ t" show "0 \ norm (flow0 s - Y s) + e / K" using e_pos K_pos by simp next show "0 < e / K" using e_pos K_pos by simp qed thus ?thesis by (simp add: algebra_simps) next assume "\0 \ t" hence "t \ 0" by simp hence [simp]: "\t\ = -t" by simp have t0_t_in_J: "{t..0} \ J" using \t \ J\ \0 \ J\ J_ivl \\ 0 \ t\ atMostAtLeast_subset_convex is_interval_convex_1 by auto note F_G_flow_cont[continuous_intros] = continuous_on_subset[OF F.flow_continuous_on] continuous_on_subset[OF G.flow_continuous_on] have "?u t + e/K \ e/K * exp(- K * t)" proof(rule gronwall_left[where g="\t. ?u t + e/K", OF _ _ _ _ K_pos order.refl \t \ 0\]) fix s assume "t \ s" "s \ 0" hence "{s..0} \ J" using t0_t_in_J by auto hence t0_s_in_existence: "{s..0} \ F.existence_ivl 0 x0" "{s..0} \ G.existence_ivl 0 x0" using J_subset G.ivl_subset_existence_ivl'[OF tG] \s \ 0\ \t \ s\ by auto hence s_in_existence: "s \ F.existence_ivl 0 x0" "s \ G.existence_ivl 0 x0" using \s \ 0\ by auto note cont_statements[continuous_intros] = F_iv_defined F.flow_in_domain G.flow_in_domain F.mem_existence_ivl_subset G.mem_existence_ivl_subset then have [continuous_intros]: "{s..0} \ T1" "{s..0} \ T2" "F.flow 0 x0 ` {s..0} \ X" "G.flow 0 x0 ` {s..0} \ X" "s \ x \ x \ 0 \ x \ F.existence_ivl 0 x0" "s \ x \ x \ 0 \ x \ G.existence_ivl 0 x0" for x using t0_s_in_existence by auto have "flow0 s - Y s = - integral {s..0} (\s. F s (flow0 s) - G s (Y s))" using t0_s_in_existence \s \ 0\ by (simp add: flow0_def Y_def ivl_integral_def F.flow_fixed_point[OF s_in_existence(1)] G.flow_fixed_point[OF s_in_existence(2)] continuous_intros integrable_on_simps Henstock_Kurzweil_Integration.integral_diff) also have "... = - integral {s..0} (\s. (F s (flow0 s) - F s (Y s)) + (F s (Y s) - G s (Y s)))" by simp also have "... = - (integral {s..0} (\s. F s (flow0 s) - F s (Y s)) + integral {s..0} (\s. F s (Y s) - G s (Y s)))" using t0_s_in_existence by (subst Henstock_Kurzweil_Integration.integral_add) (simp_all add: integral_add flow0_def Y_def continuous_intros integrable_on_simps) finally have "?u s \ norm (integral {s..0} (\s. F s (flow0 s) - F s (Y s))) + norm (integral {s..0} (\s. F s (Y s) - G s (Y s)))" by (metis (no_types, lifting) norm_minus_cancel norm_triangle_ineq) also have "... \ integral {s..0} (\s. norm (F s (flow0 s) - F s (Y s))) + integral {s..0} (\s. norm (F s (Y s) - G s (Y s)))" using t0_s_in_existence by (auto simp add: flow0_def Y_def intro!: continuous_intros continuous_on_imp_absolutely_integrable_on add_mono) also have "... \ integral {s..0} (\s. K * ?u s) + integral {s..0} (\s. e)" proof (rule add_mono[OF integral_le integral_le]) show "norm (F x (flow0 x) - F x (Y x)) \ K * norm (flow0 x - Y x)" if "x\{s..0}" for x using F_lipschitz[unfolded lipschitz_on_def, THEN conjunct2] cont_statements(1,2,4) that t0_s_in_existence F_iv_defined (* G.iv_defined *) by (metis F_lipschitz flow0_def Y_def \{s..0} \ J\ lipschitz_on_normD F.flow_in_domain G.flow_in_domain subsetCE) show "\x. x \ {s..0} \ norm (F x (Y x) - G x (Y x)) \ e" using F_G_norm_ineq Y_def \{s..0} \ J\ cont_statements(5) subset_iff t0_s_in_existence(2) using Y_def \{s..0} \ J\ cont_statements(5) subset_iff G.flow_in_domain by (metis eucl_less_le_not_le) qed (simp_all add: t0_s_in_existence continuous_intros integrable_on_simps flow0_def Y_def) also have "... = K * integral {s..0} (\s. ?u s + e / K)" using K_pos t0_s_in_existence by (simp_all add: algebra_simps Henstock_Kurzweil_Integration.integral_add t0_s_in_existence continuous_intros integrable_on_simps flow0_def Y_def) finally show "?u s + e / K \ e / K + K * integral {s..0} (\s. ?u s + e / K)" by simp next show "continuous_on {t..0} (\t. norm (flow0 t - Y t) + e / K)" using t0_t_in_J J_subset G.ivl_subset_existence_ivl'[OF tG] F_iv_defined by (auto simp add: flow0_def Y_def intro!: continuous_intros) next fix s assume "t \ s" "s \ 0" show "0 \ norm (flow0 s - Y s) + e / K" using e_pos K_pos by simp next show "0 < e / K" using e_pos K_pos by simp qed thus ?thesis by (simp add: algebra_simps) qed qed end end locale auto_ll_on_open = fixes f::"'a::{banach, heine_borel} \ 'a" and X assumes auto_local_lipschitz: "local_lipschitz UNIV X (\_::real. f)" assumes auto_open_domain[intro!, simp]: "open X" begin text \autonomous flow and existence interval \ definition "flow0 x0 t = ll_on_open.flow UNIV (\_. f) X 0 x0 t" definition "existence_ivl0 x0 = ll_on_open.existence_ivl UNIV (\_. f) X 0 x0" sublocale ll_on_open_it UNIV "\_. f" X 0 rewrites "flow = (\t0 x0 t. flow0 x0 (t - t0))" and "existence_ivl = (\t0 x0. (+) t0 ` existence_ivl0 x0)" and "(+) 0 = (\x::real. x)" and "s - 0 = s" and "(\x. x) ` S = S" and "s \ (+) t ` S \ s - t \ (S::real set)" and "P (s + t - s) = P (t::real)"\ \TODO: why does just the equation not work?\ and "P (t + s - s) = P t"\ \TODO: why does just the equation not work?\ proof - interpret ll_on_open UNIV "\_. f" X by unfold_locales (auto intro!: continuous_on_const auto_local_lipschitz) show "ll_on_open_it UNIV (\_. f) X" .. show "(+) 0 = (\x::real. x)" "(\x. x) ` S = S" "s - 0 = s" "P (t + s - s) = P t" "P (s + t - s) = P (t::real)" by auto show "flow = (\t0 x0 t. flow0 x0 (t - t0))" unfolding flow0_def - apply (rule ext) - apply (rule ext) - apply (rule flow_eq_in_existence_ivlI) - apply (auto intro: flow_shift_autonomous1 - mem_existence_ivl_shift_autonomous1 mem_existence_ivl_shift_autonomous2) - done + by (metis flow_def flow_shift_autonomous1 flow_shift_autonomous2 mem_existence_ivl_iv_defined(2)) show "existence_ivl = (\t0 x0. (+) t0 ` existence_ivl0 x0)" unfolding existence_ivl0_def apply (safe intro!: ext) subgoal using image_iff mem_existence_ivl_shift_autonomous1 by fastforce subgoal premises prems for t0 x0 x s proof - have f2: "\x1 x2. (x2::real) - x1 = - 1 * x1 + x2" by auto have "- 1 * t0 + (t0 + s) = s" by auto then show ?thesis using f2 prems mem_existence_ivl_iv_defined(2) mem_existence_ivl_shift_autonomous2 by presburger qed done show "(s \ (+) t ` S) = (s - t \ S)" by force qed \ \at this point, there should be no theorems about \existence_ivl\, only \existence_ivl0\. Moreover, \(+) _ ` _\ and \_ + _ - _\ etc should have been removed\ lemma existence_ivl_zero: "x0 \ X \ 0 \ existence_ivl0 x0" by simp lemmas [continuous_intros del] = continuous_on_f lemmas continuous_on_f_comp[continuous_intros] = continuous_on_f[OF continuous_on_const _ subset_UNIV] lemma flow_in_compact_right_existence: assumes "\t. 0 \ t \ t \ existence_ivl0 x \ flow0 x t \ K" assumes "compact K" "K \ X" assumes "x \ X" "t \ 0" shows "t \ existence_ivl0 x" proof (rule ccontr) assume "t \ existence_ivl0 x" have "bdd_above (existence_ivl0 x)" by (rule bdd_above_is_intervalI[OF is_interval_existence_ivl \0 \ t\ existence_ivl_zero]) fact+ from sup_existence_maximal[OF UNIV_I \x \ X\ assms(1-3) this] show False by auto qed lemma flow_in_compact_left_existence: assumes "\t. t \ 0 \ t \ existence_ivl0 x \ flow0 x t \ K" assumes "compact K" "K \ X" assumes "x \ X" "t \ 0" shows "t \ existence_ivl0 x" proof (rule ccontr) assume "t \ existence_ivl0 x" have "bdd_below (existence_ivl0 x)" by (rule bdd_below_is_intervalI[OF is_interval_existence_ivl \t \ 0\ _ existence_ivl_zero]) fact+ from inf_existence_minimal[OF UNIV_I \x \ X\ assms(1-3) this] show False by auto qed end locale compact_continuously_diff = derivative_on_prod T X f "\(t, x). f' x o\<^sub>L snd_blinfun" for T X and f::"real \ 'a::{banach,perfect_space,heine_borel} \ 'a" and f'::"'a \ ('a, 'a) blinfun" + assumes compact_domain: "compact X" assumes convex: "convex X" assumes nonempty_domains: "T \ {}" "X \ {}" assumes continuous_derivative: "continuous_on X f'" begin lemma ex_onorm_bound: "\B. \x \ X. norm (f' x) \ B" proof - from _ compact_domain have "compact (f' ` X)" by (intro compact_continuous_image continuous_derivative) hence "bounded (f' ` X)" by (rule compact_imp_bounded) thus ?thesis by (auto simp add: bounded_iff cball_def norm_blinfun.rep_eq) qed definition "onorm_bound = (SOME B. \x \ X. norm (f' x) \ B)" lemma onorm_bound: assumes "x \ X" shows "norm (f' x) \ onorm_bound" unfolding onorm_bound_def using someI_ex[OF ex_onorm_bound] assms by blast sublocale closed_domain X using compact_domain by unfold_locales (rule compact_imp_closed) sublocale global_lipschitz T X f onorm_bound proof (unfold_locales, rule lipschitz_onI) fix t z y assume "t \ T" "y \ X" "z \ X" then have "norm (f t y - f t z) \ onorm_bound * norm (y - z)" using onorm_bound by (intro differentiable_bound[where f'=f', OF convex]) (auto intro!: derivative_eq_intros simp: norm_blinfun.rep_eq) thus "dist (f t y) (f t z) \ onorm_bound * dist y z" by (auto simp: dist_norm norm_Pair) next from nonempty_domains obtain x where x: "x \ X" by auto show "0 \ onorm_bound" using dual_order.trans local.onorm_bound norm_ge_zero x by blast qed end \ \@{thm compact_domain}\ locale unique_on_compact_continuously_diff = self_mapping + compact_interval T + compact_continuously_diff T X f begin sublocale unique_on_closed t0 T x0 f X onorm_bound by unfold_locales (auto intro!: f' has_derivative_continuous_on) end locale c1_on_open = fixes f::"'a::{banach, perfect_space, heine_borel} \ 'a" and f' X assumes open_dom[simp]: "open X" assumes derivative_rhs: "\x. x \ X \ (f has_derivative blinfun_apply (f' x)) (at x)" assumes continuous_derivative: "continuous_on X f'" begin lemmas continuous_derivative_comp[continuous_intros] = continuous_on_compose2[OF continuous_derivative] lemma derivative_tendsto[tendsto_intros]: assumes [tendsto_intros]: "(g \ l) F" and "l \ X" shows "((\x. f' (g x)) \ f' l) F" using continuous_derivative[simplified continuous_on] assms by (auto simp: at_within_open[OF _ open_dom] intro!: tendsto_eq_intros intro: tendsto_compose) lemma c1_on_open_rev[intro, simp]: "c1_on_open (-f) (-f') X" using derivative_rhs continuous_derivative by unfold_locales (auto intro!: continuous_intros derivative_eq_intros simp: fun_Compl_def blinfun.bilinear_simps) lemma derivative_rhs_compose[derivative_intros]: "((g has_derivative g') (at x within s)) \ g x \ X \ ((\x. f (g x)) has_derivative (\xa. blinfun_apply (f' (g x)) (g' xa))) (at x within s)" by (metis has_derivative_compose[of g g' x s f "f' (g x)"] derivative_rhs) sublocale auto_ll_on_open proof (standard, rule local_lipschitzI) fix x and t::real assume "x \ X" with open_contains_cball[of "UNIV::real set"] open_UNIV open_contains_cball[of X] open_dom obtain u v where uv: "cball t u \ UNIV" "cball x v \ X" "u > 0" "v > 0" by blast let ?T = "cball t u" and ?X = "cball x v" have "bounded ?X" by simp have "compact (cball x v)" by simp interpret compact_continuously_diff ?T ?X "\_. f" f' using uv by unfold_locales (auto simp: convex_cball cball_eq_empty split_beta' intro!: derivative_eq_intros continuous_on_compose2[OF continuous_derivative] continuous_intros) have "onorm_bound-lipschitz_on ?X f" using lipschitz[of t] uv by auto thus "\u>0. \L. \t \ cball t u \ UNIV. L-lipschitz_on (cball x u \ X) f" by (intro exI[where x=v]) (auto intro!: exI[where x=onorm_bound] \0 < v\ simp: Int_absorb2 uv) qed (auto intro!: continuous_intros) end \ \@{thm derivative_rhs}\ locale c1_on_open_euclidean = c1_on_open f f' X for f::"'a::euclidean_space \ _" and f' X begin lemma c1_on_open_euclidean_anchor: True .. definition "vareq x0 t = f' (flow0 x0 t)" interpretation var: ll_on_open "existence_ivl0 x0" "vareq x0" UNIV apply standard apply (auto intro!: c1_implies_local_lipschitz[where f' = "\(t, x). vareq x0 t"] continuous_intros derivative_eq_intros simp: split_beta' blinfun.bilinear_simps vareq_def) using local.mem_existence_ivl_iv_defined(2) apply blast using local.existence_ivl_reverse local.mem_existence_ivl_iv_defined(2) apply blast using local.mem_existence_ivl_iv_defined(2) apply blast using local.existence_ivl_reverse local.mem_existence_ivl_iv_defined(2) apply blast done context begin lemma continuous_on_A[continuous_intros]: assumes "continuous_on S a" assumes "continuous_on S b" assumes "\s. s \ S \ a s \ X" assumes "\s. s \ S \ b s \ existence_ivl0 (a s)" shows "continuous_on S (\s. vareq (a s) (b s))" proof - have "continuous_on S (\x. f' (flow0 (a x) (b x)))" by (auto intro!: continuous_intros assms flow_in_domain) then show ?thesis by (rule continuous_on_eq) (auto simp: assms vareq_def) qed lemmas [intro] = mem_existence_ivl_iv_defined context fixes x0::'a begin lemma flow0_defined: "xa \ existence_ivl0 x0 \ flow0 x0 xa \ X" by (auto simp: flow_in_domain) lemma continuous_on_flow0: "continuous_on (existence_ivl0 x0) (flow0 x0)" by (auto simp: intro!: continuous_intros) lemmas continuous_on_flow0_comp[continuous_intros] = continuous_on_compose2[OF continuous_on_flow0] lemma varexivl_eq_exivl: assumes "t \ existence_ivl0 x0" shows "var.existence_ivl x0 t a = existence_ivl0 x0" proof (rule var.existence_ivl_eq_domain) fix s t x assume s: "s \ existence_ivl0 x0" and t: "t \ existence_ivl0 x0" then have "{s .. t} \ existence_ivl0 x0" by (metis atLeastatMost_empty_iff2 empty_subsetI real_Icc_closed_segment var.closed_segment_subset_domain) then have "continuous_on {s .. t} (vareq x0)" by (auto simp: closed_segment_eq_real_ivl intro!: continuous_intros flow0_defined) then have "compact ((vareq x0) ` {s .. t})" using compact_Icc by (rule compact_continuous_image) then obtain B where B: "\u. u \ {s .. t} \ norm (vareq x0 u) \ B" by (force dest!: compact_imp_bounded simp: bounded_iff) show "\M L. \t\{s..t}. \x\UNIV. norm (blinfun_apply (vareq x0 t) x) \ M + L * norm x" by (rule exI[where x=0], rule exI[where x=B]) (auto intro!: order_trans[OF norm_blinfun] mult_right_mono B simp:) qed (auto intro: assms) definition "vector_Dflow u0 t \ var.flow x0 0 u0 t" qualified abbreviation "Y z t \ flow0 (x0 + z) t" text \Linearity of the solution to the variational equation. TODO: generalize this and some other things for arbitrary linear ODEs\ lemma vector_Dflow_linear: assumes "t \ existence_ivl0 x0" shows "vector_Dflow (\ *\<^sub>R a + \ *\<^sub>R b) t = \ *\<^sub>R vector_Dflow a t + \ *\<^sub>R vector_Dflow b t" proof - note mem_existence_ivl_iv_defined[OF assms, intro, simp] have "((\c. \ *\<^sub>R var.flow x0 0 a c + \ *\<^sub>R var.flow x0 0 b c) solves_ode (\x. vareq x0 x)) (existence_ivl0 x0) UNIV" by (auto intro!: derivative_intros var.flow_has_vector_derivative solves_odeI simp: blinfun.bilinear_simps varexivl_eq_exivl vareq_def[symmetric]) moreover have "\ *\<^sub>R var.flow x0 0 a 0 + \ *\<^sub>R var.flow x0 0 b 0 = \ *\<^sub>R a + \ *\<^sub>R b" by simp moreover note is_interval_existence_ivl[of x0] ultimately show ?thesis unfolding vareq_def[symmetric] vector_Dflow_def by (rule var.maximal_existence_flow) (auto simp: assms) qed lemma linear_vector_Dflow: assumes "t \ existence_ivl0 x0" shows "linear (\z. vector_Dflow z t)" using vector_Dflow_linear[OF assms, of 1 _ 1] vector_Dflow_linear[OF assms, of _ _ 0] by (auto intro!: linearI) lemma bounded_linear_vector_Dflow: assumes "t \ existence_ivl0 x0" shows "bounded_linear (\z. vector_Dflow z t)" by (simp add: linear_linear linear_vector_Dflow assms) lemma vector_Dflow_continuous_on_time: "x0 \ X \ continuous_on (existence_ivl0 x0) (\t. vector_Dflow z t)" using var.flow_continuous_on[of x0 0 z] varexivl_eq_exivl unfolding vector_Dflow_def by (auto simp: ) proposition proposition_17_6_weak: \ \from "Differential Equations, Dynamical Systems, and an Introduction to Chaos", Hirsch/Smale/Devaney\ assumes "t \ existence_ivl0 x0" shows "(\y. (Y (y - x0) t - flow0 x0 t - vector_Dflow (y - x0) t) /\<^sub>R norm (y - x0)) \ x0 \ 0" proof- note x0_def = mem_existence_ivl_iv_defined[OF assms] have "0 \ existence_ivl0 x0" by (simp add: x0_def) text \Find some \J \ existence_ivl0 x0\ with \0 \ J\ and \t \ J\.\ define t0 where "t0 \ min 0 t" define t1 where "t1 \ max 0 t" define J where "J \ {t0..t1}" have "t0 \ 0" "0 \ t1" "0 \ J" "J \ {}" "t \ J" "compact J" and J_in_existence: "J \ existence_ivl0 x0" using ivl_subset_existence_ivl ivl_subset_existence_ivl' x0_def assms by (auto simp add: J_def t0_def t1_def min_def max_def) { fix z S assume assms: "x0 + z \ X" "S \ existence_ivl0 (x0 + z)" have "continuous_on S (Y z)" using flow_continuous_on assms(1) by (intro continuous_on_subset[OF _ assms(2)]) (simp add:) } note [continuous_intros] = this integrable_continuous_real blinfun.continuous_on have U_continuous[continuous_intros]: "\z. continuous_on J (vector_Dflow z)" by(rule continuous_on_subset[OF vector_Dflow_continuous_on_time[OF \x0 \ X\] J_in_existence]) from \t \ J\ have "t0 \ t" and "t \ t1" and "t0 \ t1" and "t0 \ existence_ivl0 x0" and "t \ existence_ivl0 x0" and "t1 \ existence_ivl0 x0" using J_def J_in_existence by auto from global_existence_ivl_explicit[OF \t0 \ existence_ivl0 x0\ \t1 \ existence_ivl0 x0\ \t0 \ t1\] obtain u K where uK_def: "0 < u" "0 < K" "ball x0 u \ X" "\y. y \ ball x0 u \ t0 \ existence_ivl0 y" "\y. y \ ball x0 u \ t1 \ existence_ivl0 y" "\t y. y \ ball x0 u \ t \ J \ dist (flow0 x0 t) (Y (y - x0) t) \ dist x0 y * exp (K * \t\)" by (auto simp add: J_def) have J_in_existence_ivl: "\y. y \ ball x0 u \ J \ existence_ivl0 y" unfolding J_def using uK_def by (simp add: real_Icc_closed_segment segment_subset_existence_ivl t0_def t1_def) have ball_in_X: "\z. z \ ball 0 u \ x0 + z \ X" using uK_def(3) by (auto simp: dist_norm) have flow0_J_props: "flow0 x0 ` J \ {}" "compact (flow0 x0 ` J)" "flow0 x0` J \ X" using \t0 \ t1\ using J_def(1) J_in_existence by (auto simp add: J_def intro!: compact_continuous_image continuous_intros flow_in_domain) have [continuous_intros]: "continuous_on J (\s. f' (flow0 x0 s))" using J_in_existence by (auto intro!: continuous_intros flow_in_domain simp:) text \ Show the thesis via cases \t = 0\, \0 < t\ and \t < 0\. \ show ?thesis proof(cases "t = 0") assume "t = 0" show ?thesis unfolding \t = 0\ Lim_at proof(simp add: dist_norm[of _ 0] del: zero_less_dist_iff, safe, rule exI, rule conjI[OF \0 < u\], safe) fix e::real and x assume "0 < e" "dist x x0 < u" hence "x \ X" using uK_def(3) by (auto simp: dist_commute) hence "inverse (norm (x - x0)) * norm (Y (x - x0) 0 - flow0 x0 0 - vector_Dflow (x - x0) 0) = 0" using x0_def by (simp add: vector_Dflow_def) thus "inverse (norm (x - x0)) * norm (flow0 x 0 - flow0 x0 0 - vector_Dflow (x - x0) 0) < e" using \0 < e\ by auto qed next assume "t \ 0" show ?thesis proof(unfold Lim_at, safe) fix e::real assume "0 < e" then obtain e' where "0 < e'" "e' < e" using dense by auto obtain N where N_ge_SupS: "Sup { norm (f' (flow0 x0 s)) |s. s \ J } \ N" (is "Sup ?S \ N") and N_gr_0: "0 < N" \ \We need N to be an upper bound of @{term ?S}, but also larger than zero.\ by (meson le_cases less_le_trans linordered_field_no_ub) have N_ineq: "\s. s \ J \ norm (f' (flow0 x0 s)) \ N" proof- fix s assume "s \ J" have "?S = (norm o f' o flow0 x0) ` J" by auto moreover have "continuous_on J (norm o f' o flow0 x0)" using J_in_existence by (auto intro!: continuous_intros) ultimately have "\a b. ?S = {a..b} \ a \ b" using continuous_image_closed_interval[OF \t0 \ t1\] by (simp add: J_def) then obtain a b where "?S = {a..b}" and "a \ b" by auto hence "bdd_above ?S" by simp from \s \ J\ cSup_upper[OF _ this] have "norm (f' (flow0 x0 s)) \ Sup ?S" by auto thus "norm (f' (flow0 x0 s)) \ N" using N_ge_SupS by simp qed text \ Define a small region around \flow0 ` J\, that is a subset of the domain \X\. \ from compact_in_open_separated[OF flow0_J_props(1,2) auto_open_domain flow0_J_props(3)] obtain e_domain where e_domain_def: "0 < e_domain" "{x. infdist x (flow0 x0 ` J) \ e_domain} \ X" by auto define G where "G \ {x\X. infdist x (flow0 x0 ` J) < e_domain}" have G_vimage: "G = ((\x. infdist x (flow0 x0 ` J)) -` {.. X" by (auto simp: G_def) have "open G" "G \ X" unfolding G_vimage by (auto intro!: open_Int open_vimage continuous_intros continuous_at_imp_continuous_on) text \Define a compact subset H of G. Inside H, we can guarantee an upper bound on the Taylor remainder.\ define e_domain2 where "e_domain2 \ e_domain / 2" have "e_domain2 > 0" "e_domain2 < e_domain" using \e_domain > 0\ by (simp_all add: e_domain2_def) define H where "H \ {x. infdist x (flow0 x0 ` J) \ e_domain2}" have H_props: "H \ {}" "compact H" "H \ G" proof- have "x0 \ flow0 x0 ` J" unfolding image_iff using \0 \ J\ x0_def by force hence "x0 \ H" using \0 < e_domain2\ by (simp add: H_def x0_def) thus "H \ {}" by auto next show "compact H" unfolding H_def using \0 < e_domain2\ flow0_J_props by (intro compact_infdist_le) simp_all next show "H \ G" proof fix x assume "x \ H" then have *: "infdist x (flow0 x0 ` J) < e_domain" using \0 < e_domain\ by (simp add: H_def e_domain2_def) then have "x \ X" using e_domain_def(2) by auto with * show "x \ G" unfolding G_def by auto qed qed have f'_cont_on_G: "(\x. x \ G \ isCont f' x)" using continuous_on_interior[OF continuous_on_subset[OF continuous_derivative \G \ X\]] by (simp add: interior_open[OF \open G\]) define e1 where "e1 \ e' / (\t\ * exp (K * \t\) * exp (N * \t\))" \ \@{term e1} is the bounding term for the Taylor remainder.\ have "0 < \t\" using \t \ 0\ by simp hence "0 < e1" using \0 < e'\ by (simp add: e1_def) text \ Taylor expansion of f on set G. \ from uniform_explicit_remainder_Taylor_1[where f=f and f'=f', OF derivative_rhs[OF subsetD[OF \G \ X\]] f'_cont_on_G \open G\ H_props \0 < e1\] obtain d_Taylor R where Taylor_expansion: "0 < d_Taylor" "\x z. f z = f x + (f' x) (z - x) + R x z" "\x y. x \ H \ y \ H \ dist x y < d_Taylor \ norm (R x y) \ e1 * dist x y" "continuous_on (G \ G) (\(a, b). R a b)" by auto text \ Find d, such that solutions are always at least \min (e_domain/2) d_Taylor\ apart, i.e. always in H. This later gives us the bound on the remainder. \ have "0 < min (e_domain/2) d_Taylor" using \0 < d_Taylor\ \0 < e_domain\ by auto from uniform_limit_flow[OF \t0 \ existence_ivl0 x0\ \t1 \ existence_ivl0 x0\ \t0 \ t1\, THEN uniform_limitD, OF this, unfolded eventually_at] obtain d_ivl where d_ivl_def: "0 < d_ivl" "\x. 0 < dist x x0 \ dist x x0 < d_ivl \ (\t\J. dist (flow0 x0 t) (Y (x - x0) t) < min (e_domain / 2) d_Taylor)" by (auto simp: dist_commute J_def) define d where "d \ min u d_ivl" have "0 < d" using \0 < u\ \0 < d_ivl\ by (simp add: d_def) hence "d \ u" "d \ d_ivl" by (auto simp: d_def) text \ Therefore, any flow0 starting in \ball x0 d\ will be in G. \ have Y_in_G: "\y. y \ ball x0 d \ (\s. Y (y - x0) s) ` J \ G" proof fix x y assume assms: "y \ ball x0 d" "x \ (\s. Y (y - x0) s) ` J" show "x \ G" proof(cases) assume "y = x0" from assms(2) have "x \ flow0 x0 ` J" by (simp add: \y = x0\) thus "x \ G" using \0 < e_domain\ \flow0 x0 ` J \ X\ by (auto simp: G_def) next assume "y \ x0" hence "0 < dist y x0" by (simp add: dist_norm) from d_ivl_def(2)[OF this] \d \ d_ivl\ \0 < e_domain\ assms(1) have dist_flow0_Y: "\t. t \ J \ dist (flow0 x0 t) (Y (y - x0) t) < e_domain" by (auto simp: dist_commute) from assms(2) obtain t where t_def: "t \ J" "x = Y (y - x0) t" by auto have "x \ X" unfolding t_def(2) using uK_def(3) assms(1) \d \ u\ subsetD[OF J_in_existence_ivl t_def(1)] by (auto simp: intro!: flow_in_domain) have "flow0 x0 t \ flow0 x0 ` J" using t_def by auto from dist_flow0_Y[OF t_def(1)] have "dist x (flow0 x0 t) < e_domain" by (simp add: t_def(2) dist_commute) from le_less_trans[OF infdist_le[OF \flow0 x0 t \ flow0 x0 ` J\] this] \x \ X\ show "x \ G" by (auto simp: G_def) qed qed from this[of x0] \0 < d\ have X_in_G: "flow0 x0 ` J \ G" by simp show "\d>0. \x. 0 < dist x x0 \ dist x x0 < d \ dist ((Y (x - x0) t - flow0 x0 t - vector_Dflow (x - x0) t) /\<^sub>R norm (x - x0)) 0 < e" proof(rule exI, rule conjI[OF \0 < d\], safe, unfold norm_conv_dist[symmetric]) fix x assume x_x0_dist: "0 < dist x x0" "dist x x0 < d" hence x_in_ball': "x \ ball x0 d" by (simp add: dist_commute) hence x_in_ball: "x \ ball x0 u" using \d \ u\ by simp text \ First, some prerequisites. \ from x_in_ball have z_in_ball: "x - x0 \ ball 0 u" using \0 < u\ by (simp add: dist_norm) hence [continuous_intros]: "dist x0 x < u" by (auto simp: dist_norm) from J_in_existence_ivl[OF x_in_ball] have J_in_existence_ivl_x: "J \ existence_ivl0 x" . from ball_in_X[OF z_in_ball] have x_in_X[continuous_intros]: "x \ X" by simp text \ On all of \J\, we can find upper bounds for the distance of \flow0\ and \Y\. \ have dist_flow0_Y: "\s. s \ J \ dist (flow0 x0 s) (Y (x - x0) s) \ dist x0 x * exp (K * \t\)" using t0_def t1_def uK_def(2) by (intro order_trans[OF uK_def(6)[OF x_in_ball] mult_left_mono]) (auto simp add: J_def intro!: mult_mono) from d_ivl_def x_x0_dist \d \ d_ivl\ have dist_flow0_Y2: "\t. t \ J \ dist (flow0 x0 t) (Y (x - x0) t) < min (e_domain2) d_Taylor" by (auto simp: e_domain2_def) let ?g = "\t. norm (Y (x - x0) t - flow0 x0 t - vector_Dflow (x - x0) t)" let ?C = "\t\ * dist x0 x * exp (K * \t\) * e1" text \ Find an upper bound to \?g\, i.e. show that \?g s \ ?C + N * integral {a..b} ?g\ for \{a..b} = {0..s}\ or \{a..b} = {s..0}\ for some \s \ J\. We can then apply Grönwall's inequality to obtain a true bound for \?g\. \ have g_bound: "?g s \ ?C + N * integral {a..b} ?g" if s_def: "s \ {a..b}" and J'_def: "{a..b} \ J" and ab_cases: "(a = 0 \ b = s) \ (a = s \ b = 0)" for s a b proof - from that have "s \ J" by auto have s_in_existence_ivl_x0: "s \ existence_ivl0 x0" using J_in_existence \s \ J\ by auto have s_in_existence_ivl: "\y. y \ ball x0 u \ s \ existence_ivl0 y" using J_in_existence_ivl \s \ J\ by auto have s_in_existence_ivl2: "\z. z \ ball 0 u \ s \ existence_ivl0 (x0 + z)" using s_in_existence_ivl by (simp add: dist_norm) text \Prove continuities beforehand.\ note continuous_on_0_s[continuous_intros] = continuous_on_subset[OF _ \{a..b} \ J\] have[continuous_intros]: "continuous_on J (flow0 x0)" using J_in_existence by (auto intro!: continuous_intros simp:) { fix z S assume assms: "x0 + z \ X" "S \ existence_ivl0 (x0 + z)" have "continuous_on S (\s. f (Y z s))" proof(rule continuous_on_subset[OF _ assms(2)]) show "continuous_on (existence_ivl0 (x0 + z)) (\s. f (Y z s))" using assms by (auto intro!: continuous_intros flow_in_domain flow_continuous_on simp:) qed } note [continuous_intros] = this have [continuous_intros]: "continuous_on J (\s. f (flow0 x0 s))" by(rule continuous_on_subset[OF _ J_in_existence]) (auto intro!: continuous_intros flow_continuous_on flow_in_domain simp: x0_def) have [continuous_intros]: "\z. continuous_on J (\s. f' (flow0 x0 s) (vector_Dflow z s))" proof- fix z have a1: "continuous_on J (flow0 x0)" by (auto intro!: continuous_intros) have a2: "(\s. (flow0 x0 s, vector_Dflow z s)) ` J \ (flow0 x0 ` J) \ ((\s. vector_Dflow z s) ` J)" by auto have a3: "continuous_on ((\s. (flow0 x0 s, vector_Dflow z s)) ` J) (\(x, u). f' x u)" using assms flow0_J_props by (auto intro!: continuous_intros simp: split_beta') from continuous_on_compose[OF continuous_on_Pair[OF a1 U_continuous] a3] show "continuous_on J (\s. f' (flow0 x0 s) (vector_Dflow z s))" by simp qed have [continuous_intros]: "continuous_on J (\s. R (flow0 x0 s) (Y (x - x0) s))" using J_in_existence J_in_existence_ivl[OF x_in_ball] X_in_G \{a..b} \ J\ Y_in_G x_x0_dist by (auto intro!: continuous_intros continuous_on_compose_Pair[OF Taylor_expansion(4)] simp: dist_commute subset_iff) hence [continuous_intros]: "(\s. R (flow0 x0 s) (Y (x - x0) s)) integrable_on J" unfolding J_def by (rule integrable_continuous_real) have i1: "integral {a..b} (\s. f (flow0 x s)) - integral {a..b} (\s. f (flow0 x0 s)) = integral {a..b} (\s. f (flow0 x s) - f (flow0 x0 s))" using J_in_existence_ivl[OF x_in_ball] apply (intro Henstock_Kurzweil_Integration.integral_diff[symmetric]) by (auto intro!: continuous_intros existence_ivl_reverse) have i2: "integral {a..b} (\s. f (flow0 x s) - f (flow0 x0 s) - (f' (flow0 x0 s)) (vector_Dflow (x - x0) s)) = integral {a..b} (\s. f (flow0 x s) - f (flow0 x0 s)) - integral {a..b} (\s. f' (flow0 x0 s) (vector_Dflow (x - x0) s))" using J_in_existence_ivl[OF x_in_ball] by (intro Henstock_Kurzweil_Integration.integral_diff[OF Henstock_Kurzweil_Integration.integrable_diff]) (auto intro!: continuous_intros existence_ivl_reverse) from ab_cases have "?g s = norm (integral {a..b} (\s'. f (Y (x - x0) s')) - integral {a..b} (\s'. f (flow0 x0 s')) - integral {a..b} (\s'. (f' (flow0 x0 s')) (vector_Dflow (x - x0) s')))" proof(safe) assume "a = 0" "b = s" hence "0 \ s" using \s \ {a..b}\ by simp text\ Integral equations for flow0, Y and U. \ have flow0_integral_eq: "flow0 x0 s = x0 + ivl_integral 0 s (\s. f (flow0 x0 s))" by (rule flow_fixed_point[OF s_in_existence_ivl_x0]) have Y_integral_eq: "flow0 x s = x0 + (x - x0) + ivl_integral 0 s (\s. f (Y (x - x0) s))" using flow_fixed_point \0 \ s\ s_in_existence_ivl2[OF z_in_ball] ball_in_X[OF z_in_ball] by (simp add:) have U_integral_eq: "vector_Dflow (x - x0) s = (x - x0) + ivl_integral 0 s (\s. vareq x0 s (vector_Dflow (x - x0) s))" unfolding vector_Dflow_def by (rule var.flow_fixed_point) (auto simp: \0 \ s\ x0_def varexivl_eq_exivl s_in_existence_ivl_x0) show "?g s = norm (integral {0..s} (\s'. f (Y (x - x0) s')) - integral {0..s} (\s'. f (flow0 x0 s')) - integral {0..s} (\s'. blinfun_apply (f' (flow0 x0 s')) (vector_Dflow (x - x0) s')))" using \0 \ s\ unfolding vareq_def[symmetric] by (simp add: flow0_integral_eq Y_integral_eq U_integral_eq ivl_integral_def) next assume "a = s" "b = 0" hence "s \ 0" using \s \ {a..b}\ by simp have flow0_integral_eq_left: "flow0 x0 s = x0 + ivl_integral 0 s (\s. f (flow0 x0 s))" by (rule flow_fixed_point[OF s_in_existence_ivl_x0]) have Y_integral_eq_left: "Y (x - x0) s = x0 + (x - x0) + ivl_integral 0 s (\s. f (Y (x - x0) s))" using flow_fixed_point \s \ 0\ s_in_existence_ivl2[OF z_in_ball] ball_in_X[OF z_in_ball] by simp have U_integral_eq_left: "vector_Dflow (x - x0) s = (x - x0) + ivl_integral 0 s (\s. vareq x0 s (vector_Dflow (x - x0) s))" unfolding vector_Dflow_def by (rule var.flow_fixed_point) (auto simp: \s \ 0\ x0_def varexivl_eq_exivl s_in_existence_ivl_x0) have "?g s = norm (- integral {s..0} (\s'. f (Y (x - x0) s')) + integral {s..0} (\s'. f (flow0 x0 s')) + integral {s..0} (\s'. vareq x0 s' (vector_Dflow (x - x0) s')))" unfolding flow0_integral_eq_left Y_integral_eq_left U_integral_eq_left using \s \ 0\ by (simp add: ivl_integral_def) also have "... = norm (integral {s..0} (\s'. f (Y (x - x0) s')) - integral {s..0} (\s'. f (flow0 x0 s')) - integral {s..0} (\s'. vareq x0 s' (vector_Dflow (x - x0) s')))" by (subst norm_minus_cancel[symmetric], simp) finally show "?g s = norm (integral {s..0} (\s'. f (Y (x - x0) s')) - integral {s..0} (\s'. f (flow0 x0 s')) - integral {s..0} (\s'. blinfun_apply (f' (flow0 x0 s')) (vector_Dflow (x - x0) s')))" unfolding vareq_def . qed also have "... = norm (integral {a..b} (\s. f (Y (x - x0) s) - f (flow0 x0 s) - (f' (flow0 x0 s)) (vector_Dflow (x - x0) s)))" by (simp add: i1 i2) also have "... \ integral {a..b} (\s. norm (f (Y (x - x0) s) - f (flow0 x0 s) - f' (flow0 x0 s) (vector_Dflow (x - x0) s)))" using x_in_X J_in_existence_ivl_x J_in_existence \{a..b} \ J\ by (auto intro!: continuous_intros continuous_on_imp_absolutely_integrable_on existence_ivl_reverse) also have "... = integral {a..b} (\s. norm (f' (flow0 x0 s) (Y (x - x0) s - flow0 x0 s - vector_Dflow (x - x0) s) + R (flow0 x0 s) (Y (x - x0) s)))" proof (safe intro!: integral_spike[OF negligible_empty, simplified] arg_cong[where f=norm]) fix s' assume "s' \ {a..b}" show "f' (flow0 x0 s') (Y (x - x0) s' - flow0 x0 s' - vector_Dflow (x - x0) s') + R (flow0 x0 s') (Y (x - x0) s') = f (Y (x - x0) s') - f (flow0 x0 s') - f' (flow0 x0 s') (vector_Dflow (x - x0) s')" by (simp add: blinfun.diff_right Taylor_expansion(2)[of "flow0 x s'" "flow0 x0 s'"]) qed also have "... \ integral {a..b} (\s. norm (f' (flow0 x0 s) (Y (x - x0) s - flow0 x0 s - vector_Dflow (x - x0) s)) + norm (R (flow0 x0 s) (Y (x - x0) s)))" using J_in_existence_ivl[OF x_in_ball] norm_triangle_ineq using \continuous_on J (\s. R (flow0 x0 s) (Y (x - x0) s))\ by (auto intro!: continuous_intros integral_le) also have "... = integral {a..b} (\s. norm (f' (flow0 x0 s) (Y (x - x0) s - flow0 x0 s - vector_Dflow (x - x0) s))) + integral {a..b} (\s. norm (R (flow0 x0 s) (Y (x - x0) s)))" using J_in_existence_ivl[OF x_in_ball] using \continuous_on J (\s. R (flow0 x0 s) (Y (x - x0) s))\ by (auto intro!: continuous_intros Henstock_Kurzweil_Integration.integral_add) also have "... \ N * integral {a..b} ?g + ?C" (is "?l1 + ?r1 \ _") proof(rule add_mono) have "?l1 \ integral {a..b} (\s. norm (f' (flow0 x0 s)) * norm (Y (x - x0) s - flow0 x0 s - vector_Dflow (x - x0) s))" using norm_blinfun J_in_existence_ivl[OF x_in_ball] by (auto intro!: continuous_intros integral_le) also have "... \ integral {a..b} (\s. N * norm (Y (x - x0) s - flow0 x0 s - vector_Dflow (x - x0) s))" using J_in_existence_ivl[OF x_in_ball] N_ineq[OF \{a..b} \ J\[THEN subsetD]] by (intro integral_le) (auto intro!: continuous_intros mult_right_mono) also have "... = N * integral {a..b} (\s. norm ((Y (x - x0) s - flow0 x0 s - vector_Dflow (x - x0) s)))" unfolding real_scaleR_def[symmetric] by(rule integral_cmul) finally show "?l1 \ N * integral {a..b} ?g" . next have "?r1 \ integral {a..b} (\s. e1 * dist (flow0 x0 s) (Y (x - x0) s))" using J_in_existence_ivl[OF x_in_ball] \0 < e_domain\ dist_flow0_Y2 \0 < e_domain2\ by (intro integral_le) (force intro!: continuous_intros Taylor_expansion(3) order_trans[OF infdist_le] dest!: \{a..b} \ J\[THEN subsetD] intro: less_imp_le simp: dist_commute H_def)+ also have "... \ integral {a..b} (\s. e1 * (dist x0 x * exp (K * \t\)))" apply(rule integral_le) subgoal using J_in_existence_ivl[OF x_in_ball] by (force intro!: continuous_intros) subgoal by force subgoal by (force dest!: \{a..b} \ J\[THEN subsetD] intro!: less_imp_le[OF \0 < e1\] mult_left_mono[OF dist_flow0_Y]) done also have "... \ ?C" using \s \ J\ x_x0_dist \0 < e1\ \{a..b} \ J\ \0 < \t\\ t0_def t1_def by (auto simp: integral_const_real J_def(1)) finally show "?r1 \ ?C" . qed finally show ?thesis by simp qed have g_continuous: "continuous_on J ?g" using J_in_existence_ivl[OF x_in_ball] J_in_existence using J_def(1) U_continuous by (auto simp: J_def intro!: continuous_intros) note [continuous_intros] = continuous_on_subset[OF g_continuous] have C_gr_zero: "0 < ?C" using \0 < \t\\ \0 < e1\ x_x0_dist(1) by (simp add: dist_commute) have "0 \ t \ t \ 0" by auto then have "?g t \ ?C * exp (N * \t\)" proof assume "0 \ t" moreover have "continuous_on {0..t} (vector_Dflow (x - x0))" using U_continuous by (rule continuous_on_subset) (auto simp: J_def t0_def t1_def) then have "norm (Y (x - x0) t - flow0 x0 t - vector_Dflow (x - x0) t) \ \t\ * dist x0 x * exp (K * \t\) * e1 * exp (N * t)" using \t \ J\ J_def \t0 \ 0\ J_in_existence J_in_existence_ivl_x by (intro gronwall[OF g_bound _ _ C_gr_zero \0 < N\ \0 \ t\ order.refl]) (auto intro!: continuous_intros simp: ) ultimately show ?thesis by simp next assume "t \ 0" moreover have "continuous_on {t .. 0} (vector_Dflow (x - x0))" using U_continuous by (rule continuous_on_subset) (auto simp: J_def t0_def t1_def) then have "norm (Y (x - x0) t - flow0 x0 t - vector_Dflow (x - x0) t) \ \t\ * dist x0 x * exp (K * \t\) * e1 * exp (- N * t)" using \t \ J\ J_def \0 \ t1\ J_in_existence J_in_existence_ivl_x by (intro gronwall_left[OF g_bound _ _ C_gr_zero \0 < N\ order.refl \t \ 0\]) (auto intro!: continuous_intros) ultimately show ?thesis by simp qed also have "... = dist x x0 * (\t\ * exp (K * \t\) * e1 * exp (N * \t\))" by (auto simp: dist_commute) also have "... < norm (x - x0) * e" unfolding e1_def using \e' < e\ \0 < \t\\ \0 < e1\ x_x0_dist(1) by (simp add: dist_norm) finally show "norm ((Y (x - x0) t - flow0 x0 t - vector_Dflow (x - x0) t) /\<^sub>R norm (x - x0)) < e" by (simp, metis x_x0_dist(1) dist_norm divide_inverse mult.commute pos_divide_less_eq) qed qed qed qed lemma local_lipschitz_A: "OT \ existence_ivl0 x0 \ local_lipschitz OT (OS::('a \\<^sub>L 'a) set) (\t. (o\<^sub>L) (vareq x0 t))" by (rule local_lipschitz_subset[OF _ _ subset_UNIV, where T="existence_ivl0 x0"]) (auto simp: split_beta' vareq_def intro!: c1_implies_local_lipschitz[where f'="\(t, x). comp3 (f' (flow0 x0 t))"] derivative_eq_intros blinfun_eqI ext continuous_intros flow_in_domain) lemma total_derivative_ll_on_open: "ll_on_open (existence_ivl0 x0) (\t. blinfun_compose (vareq x0 t)) (UNIV::('a \\<^sub>L 'a) set)" by standard (auto intro!: continuous_intros local_lipschitz_A[OF order_refl]) end end sublocale mvar: ll_on_open "existence_ivl0 x0" "\t. blinfun_compose (vareq x0 t)" "UNIV::('a \\<^sub>L 'a) set" for x0 by (rule total_derivative_ll_on_open) lemma mvar_existence_ivl_eq_existence_ivl[simp]:\ \TODO: unify with @{thm varexivl_eq_exivl}\ assumes "t \ existence_ivl0 x0" shows "mvar.existence_ivl x0 t = (\_. existence_ivl0 x0)" proof (rule ext, rule mvar.existence_ivl_eq_domain) fix s t x assume s: "s \ existence_ivl0 x0" and t: "t \ existence_ivl0 x0" then have "{s .. t} \ existence_ivl0 x0" by (meson atLeastAtMost_iff is_interval_1 is_interval_existence_ivl subsetI) then have "continuous_on {s .. t} (vareq x0)" by (auto intro!: continuous_intros) then have "compact (vareq x0 ` {s .. t})" using compact_Icc by (rule compact_continuous_image) then obtain B where B: "\u. u \ {s .. t} \ norm (vareq x0 u) \ B" by (force dest!: compact_imp_bounded simp: bounded_iff) show "\M L. \t\{s .. t}. \x\UNIV. norm (vareq x0 t o\<^sub>L x) \ M + L * norm x" unfolding o_def by (rule exI[where x=0], rule exI[where x=B]) (auto intro!: order_trans[OF norm_blinfun_compose] mult_right_mono B) qed (auto intro: assms) lemma assumes "t \ existence_ivl0 x0" shows "continuous_on (UNIV \ existence_ivl0 x0) (\(x, ta). mvar.flow x0 t x ta)" proof - from mvar.flow_continuous_on_state_space[of x0 t, unfolded mvar_existence_ivl_eq_existence_ivl[OF assms]] show "continuous_on (UNIV \ existence_ivl0 x0) (\(x, ta). mvar.flow x0 t x ta)" . qed definition "Dflow x0 = mvar.flow x0 0 id_blinfun" lemma var_eq_mvar: assumes "t0 \ existence_ivl0 x0" assumes "t \ existence_ivl0 x0" shows "var.flow x0 t0 i t = mvar.flow x0 t0 id_blinfun t i" by (rule var.flow_unique) (auto intro!: assms derivative_eq_intros mvar.flow_has_derivative simp: varexivl_eq_exivl assms has_vector_derivative_def blinfun.bilinear_simps) lemma Dflow_zero[simp]: "x \ X \ Dflow x 0 = 1\<^sub>L" unfolding Dflow_def by (subst mvar.flow_initial_time) auto subsection \Differentiability of the flow0\ text \ \U t\, i.e. the solution of the variational equation, is the space derivative at the initial value \x0\. \ lemma flow_dx_derivative: assumes "t \ existence_ivl0 x0" shows "((\x0. flow0 x0 t) has_derivative (\z. vector_Dflow x0 z t)) (at x0)" unfolding has_derivative_at2 using assms by (intro iffD1[OF LIM_equal proposition_17_6_weak[OF assms]] conjI[OF bounded_linear_vector_Dflow[OF assms]]) (simp add: diff_diff_add inverse_eq_divide) lemma flow_dx_derivative_blinfun: assumes "t \ existence_ivl0 x0" shows "((\x. flow0 x t) has_derivative Blinfun (\z. vector_Dflow x0 z t)) (at x0)" by (rule has_derivative_Blinfun[OF flow_dx_derivative[OF assms]]) definition "flowderiv x0 t = comp12 (Dflow x0 t) (blinfun_scaleR_left (f (flow0 x0 t)))" lemma flowderiv_eq: "flowderiv x0 t (\\<^sub>1, \\<^sub>2) = (Dflow x0 t) \\<^sub>1 + \\<^sub>2 *\<^sub>R f (flow0 x0 t)" by (auto simp: flowderiv_def) lemma W_continuous_on: "continuous_on (Sigma X existence_ivl0) (\(x0, t). Dflow x0 t)" \ \TODO: somewhere here is hidden continuity wrt rhs of ODE, extract it!\ unfolding continuous_on split_beta' proof (safe intro!: tendstoI) fix e'::real and t x assume x: "x \ X" and tx: "t \ existence_ivl0 x" and e': "e' > 0" let ?S = "Sigma X existence_ivl0" have "(x, t) \ ?S" using x tx by auto from open_prod_elim[OF open_state_space this] obtain OX OT where OXOT: "open OX" "open OT" "(x, t) \ OX \ OT" "OX \ OT \ ?S" by blast then obtain dx dt where dx: "dx > 0" "cball x dx \ OX" and dt: "dt > 0" "cball t dt \ OT" by (force simp: open_contains_cball) from OXOT dt dx have "cball t dt \ existence_ivl0 x" "cball x dx \ X" apply (auto simp: subset_iff) subgoal for ta apply (drule spec[where x=ta]) apply (drule spec[where x=t])+ apply auto done done have one_exivl: "mvar.existence_ivl x 0 = (\_. existence_ivl0 x)" by (rule mvar_existence_ivl_eq_existence_ivl[OF existence_ivl_zero[OF \x \ X\]]) have *: "closed ({t .. 0} \ {0 .. t})" "{t .. 0} \ {0 .. t} \ {}" by auto let ?T = "{t .. 0} \ {0 .. t} \ cball t dt" have "compact ?T" by (auto intro!: compact_Un) have "?T \ existence_ivl0 x" by (intro Un_least ivl_subset_existence_ivl' ivl_subset_existence_ivl \x \ X\ \t \ existence_ivl0 x\ \cball t dt \ existence_ivl0 x\) have "compact (mvar.flow x 0 id_blinfun ` ?T)" using \?T \ _\ \x \ X\ mvar_existence_ivl_eq_existence_ivl[OF existence_ivl_zero[OF \x \ X\]] by (auto intro!: \0 < dx\ compact_continuous_image \compact ?T\ continuous_on_subset[OF mvar.flow_continuous_on]) let ?line = "mvar.flow x 0 id_blinfun ` ?T" let ?X = "{x. infdist x ?line \ dx}" have "compact ?X" using \?T \ _\ \x \ X\ mvar_existence_ivl_eq_existence_ivl[OF existence_ivl_zero[OF \x \ X\]] by (auto intro!: compact_infdist_le \0 < dx\ compact_continuous_image compact_Un continuous_on_subset[OF mvar.flow_continuous_on ]) from mvar.local_lipschitz \?T \ _\ have llc: "local_lipschitz ?T ?X (\t. (o\<^sub>L) (vareq x t))" by (rule local_lipschitz_subset) auto have cont: "\xa. xa \ ?X \ continuous_on ?T (\t. vareq x t o\<^sub>L xa)" using \?T \ _\ by (auto intro!: continuous_intros \x \ X\) from local_lipschitz_compact_implies_lipschitz[OF llc \compact ?X\ \compact ?T\ cont] obtain K' where K': "\ta. ta \ ?T \ K'-lipschitz_on ?X ((o\<^sub>L) (vareq x ta))" by blast define K where "K \ abs K' + 1" have "K > 0" by (simp add: K_def) have K: "\ta. ta \ ?T \ K-lipschitz_on ?X ((o\<^sub>L) (vareq x ta))" by (auto intro!: lipschitz_onI mult_right_mono order_trans[OF lipschitz_onD[OF K']] simp: K_def) have ex_ivlI: "\y. y \ cball x dx \ ?T \ existence_ivl0 y" using dx dt OXOT by (intro Un_least ivl_subset_existence_ivl' ivl_subset_existence_ivl; force) have cont: "continuous_on ((?T \ ?X) \ cball x dx) (\((ta, xa), y). (vareq y ta o\<^sub>L xa))" using \cball x dx \ X\ ex_ivlI by (force intro!: continuous_intros simp: split_beta' ) have "mvar.flow x 0 id_blinfun t \ mvar.flow x 0 id_blinfun ` ({t..0} \ {0..t} \ cball t dt)" by auto then have mem: "(t, mvar.flow x 0 id_blinfun t, x) \ ?T \ ?X \ cball x dx" by (auto simp: \0 < dx\ less_imp_le) define e where "e \ min e' (dx / 2) / 2" have "e > 0" using \e' > 0\ by (auto simp: e_def \0 < dx\) define d where "d \ e * K / (exp (K * (abs t + abs dt + 1)) - 1)" have "d > 0" by (auto simp: d_def intro!: mult_pos_pos divide_pos_pos \0 < e\ \K > 0\) have cmpct: "compact ((?T \ ?X) \ cball x dx)" "compact (?T \ ?X)" using \compact ?T\ \compact ?X\ by (auto intro!: compact_cball compact_Times) have compact_line: "compact ?line" using \{t..0} \ {0..t} \ cball t dt \ existence_ivl0 x\ one_exivl by (force intro!: compact_continuous_image \compact ?T\ continuous_on_subset[OF mvar.flow_continuous_on] simp: \x \ X\) from compact_uniformly_continuous[OF cont cmpct(1), unfolded uniformly_continuous_on_def, rule_format, OF \0 < d\] obtain d' where d': "d' > 0" "\ta xa xa' y. ta \ ?T \ xa \ ?X \ xa'\cball x dx \ y\cball x dx \ dist xa' y < d' \ dist (vareq xa' ta o\<^sub>L xa) (vareq y ta o\<^sub>L xa) < d" by (auto simp: dist_prod_def) { fix y assume dxy: "dist x y < d'" assume "y \ cball x dx" then have "y \ X" using dx dt OXOT by force+ have two_exivl: "mvar.existence_ivl y 0 = (\_. existence_ivl0 y)" by (rule mvar_existence_ivl_eq_existence_ivl[OF existence_ivl_zero[OF \y \ X\]]) let ?X' = "\x \ ?line. ball x dx" have "open ?X'" by auto have "?X' \ ?X" by (auto intro!: infdist_le2 simp: dist_commute) interpret oneR: ll_on_open "existence_ivl0 x" "(\t. (o\<^sub>L) (vareq x t))" ?X' by standard (auto intro!: \x \ X\ continuous_intros local_lipschitz_A[OF order_refl]) interpret twoR: ll_on_open "existence_ivl0 y" "(\t. (o\<^sub>L) (vareq y t))" ?X' by standard (auto intro!: \y \ X\ continuous_intros local_lipschitz_A[OF order_refl]) interpret both: two_ll_on_open "(\t. (o\<^sub>L) (vareq x t))" "existence_ivl0 x" "(\t. (o\<^sub>L) (vareq y t))" "existence_ivl0 y" ?X' ?T "id_blinfun" d K proof unfold_locales show "0 < K" by (simp add: \0 < K\) show iv_defined: "0 \ {t..0} \ {0..t} \ cball t dt" by auto show "is_interval ({t..0} \ {0..t} \ cball t dt)" by (auto simp: is_interval_def dist_real_def) show "{t..0} \ {0..t} \ cball t dt \ oneR.existence_ivl 0 id_blinfun" apply (rule oneR.maximal_existence_flow[where x="mvar.flow x 0 id_blinfun"]) subgoal apply (rule solves_odeI) apply (rule has_vderiv_on_subset[OF solves_odeD(1)[OF mvar.flow_solves_ode[of 0 x id_blinfun]]]) subgoal using \x \ X\ \?T \ _\ \0 < dx\ by simp subgoal by simp subgoal by (simp add: \cball t dt \ existence_ivl0 x\ ivl_subset_existence_ivl ivl_subset_existence_ivl' one_exivl tx) subgoal using dx by (auto; force) done subgoal by (simp add: \x \ X\) subgoal by fact subgoal using iv_defined by blast subgoal using \{t..0} \ {0..t} \ cball t dt \ existence_ivl0 x\ by blast done fix s assume s: "s \ ?T" then show "K-lipschitz_on ?X' ((o\<^sub>L) (vareq x s))" by (intro lipschitz_on_subset[OF K \?X' \ ?X\]) auto fix j assume j: "j \ ?X'" show "norm ((vareq x s o\<^sub>L j) - (vareq y s o\<^sub>L j)) < d" unfolding dist_norm[symmetric] apply (rule d') subgoal by (rule s) subgoal using \?X' \ ?X\ j .. subgoal using \dx > 0\ by simp subgoal using \y \ cball x dx\ by simp subgoal using dxy by simp done qed have less_e: "norm (Dflow x s - both.Y s) < e" if s: "s \ ?T \ twoR.existence_ivl 0 id_blinfun" for s proof - from s have s_less: "\s\ < \t\ + \dt\ + 1" by (auto simp: dist_real_def) note both.norm_X_Y_bound[rule_format, OF s] also have "d / K * (exp (K * \s\) - 1) = e * ((exp (K * \s\) - 1) / (exp (K * (\t\ + \dt\ + 1)) - 1))" by (simp add: d_def) also have "\ < e * 1" by (rule mult_strict_left_mono[OF _ \0 < e\]) (simp add: add_nonneg_pos \0 < K\ \0 < e\ s_less) also have "\ = e" by simp also from s have s: "s \ ?T" by simp have "both.flow0 s = Dflow x s" unfolding both.flow0_def Dflow_def apply (rule oneR.maximal_existence_flow[where K="?T"]) subgoal apply (rule solves_odeI) apply (rule has_vderiv_on_subset[OF solves_odeD(1)[OF mvar.flow_solves_ode[of 0 x id_blinfun]]]) subgoal using \x \ X\ \0 < dx\ by simp subgoal by simp subgoal by (simp add: \cball t dt \ existence_ivl0 x\ ivl_subset_existence_ivl ivl_subset_existence_ivl' one_exivl tx) subgoal using dx by (auto; force) done subgoal by (simp add: \x \ X\) subgoal by (rule both.J_ivl) subgoal using both.t0_in_J by blast subgoal using \{t..0} \ {0..t} \ cball t dt \ existence_ivl0 x\ by blast subgoal using s by blast done finally show ?thesis . qed have "e < dx" using \dx > 0\ by (auto simp: e_def) let ?i = "{y. infdist y (mvar.flow x 0 id_blinfun ` ?T) \ e}" have 1: "?i \ (\x\mvar.flow x 0 id_blinfun ` ?T. ball x dx)" proof - have cl: "closed ?line" "?line \ {}" using compact_line by (auto simp: compact_imp_closed) have "?i \ (\y\mvar.flow x 0 id_blinfun ` ?T. cball y e)" proof safe fix x assume H: "infdist x ?line \ e" from infdist_attains_inf[OF cl, of x] obtain y where "y \ ?line" "infdist x ?line = dist x y" by auto then show "x \ (\x\?line. cball x e)" using H by (auto simp: dist_commute) qed also have "\ \ (\x\?line. ball x dx)" using \e < dx\ by auto finally show ?thesis . qed have 2: "twoR.flow 0 id_blinfun s \ ?i" if "s \ ?T" "s \ twoR.existence_ivl 0 id_blinfun" for s proof - from that have sT: "s \ ?T \ twoR.existence_ivl 0 id_blinfun" by force from less_e[OF this] have "dist (twoR.flow 0 id_blinfun s) (mvar.flow x 0 id_blinfun s) \ e" unfolding Dflow_def both.Y_def dist_commute dist_norm by simp then show ?thesis using sT by (force intro: infdist_le2) qed have T_subset: "?T \ twoR.existence_ivl 0 id_blinfun" apply (rule twoR.subset_mem_compact_implies_subset_existence_interval[ where K="{x. infdist x ?line \ e}"]) subgoal using \0 < dt\ by force subgoal by (rule both.J_ivl) subgoal using \y \ cball x dx\ ex_ivlI by blast subgoal using both.F_iv_defined(2) by blast subgoal by (rule 2) subgoal using \dt > 0\ by (intro compact_infdist_le) (auto intro!: compact_line \0 < e\) subgoal by (rule 1) done also have "twoR.existence_ivl 0 id_blinfun \ existence_ivl0 y" by (rule twoR.existence_ivl_subset) finally have "?T \ existence_ivl0 y" . have "norm (Dflow x s - Dflow y s) < e" if s: "s \ ?T" for s proof - from s have "s \ ?T \ twoR.existence_ivl 0 id_blinfun" using T_subset by force from less_e[OF this] have "norm (Dflow x s - both.Y s) < e" . also have "mvar.flow y 0 id_blinfun s = twoR.flow 0 id_blinfun s" apply (rule mvar.maximal_existence_flow[where K="?T"]) subgoal apply (rule solves_odeI) apply (rule has_vderiv_on_subset[OF solves_odeD(1)[OF twoR.flow_solves_ode[of 0 id_blinfun]]]) subgoal using \y \ X\ by simp subgoal using both.F_iv_defined(2) by blast subgoal using T_subset by blast subgoal by simp done subgoal using \y \ X\ auto_ll_on_open.existence_ivl_zero auto_ll_on_open_axioms both.F_iv_defined(2) twoR.flow_initial_time by blast subgoal by (rule both.J_ivl) subgoal using both.t0_in_J by blast subgoal using \{t..0} \ {0..t} \ cball t dt \ existence_ivl0 y\ by blast subgoal using s by blast done then have "both.Y s = Dflow y s" unfolding both.Y_def Dflow_def by simp finally show ?thesis . qed } note cont_data = this have "\\<^sub>F (y, s) in at (x, t) within ?S. dist x y < d'" unfolding at_within_open[OF \(x, t) \ ?S\ open_state_space] UNIV_Times_UNIV[symmetric] using \d' > 0\ by (intro eventually_at_Pair_within_TimesI1) (auto simp: eventually_at less_imp_le dist_commute) moreover have "\\<^sub>F (y, s) in at (x, t) within ?S. y \ cball x dx" unfolding at_within_open[OF \(x, t) \ ?S\ open_state_space] UNIV_Times_UNIV[symmetric] using \dx > 0\ by (intro eventually_at_Pair_within_TimesI1) (auto simp: eventually_at less_imp_le dist_commute) moreover have "\\<^sub>F (y, s) in at (x, t) within ?S. s \ ?T" unfolding at_within_open[OF \(x, t) \ ?S\ open_state_space] UNIV_Times_UNIV[symmetric] using \dt > 0\ by (intro eventually_at_Pair_within_TimesI2) (auto simp: eventually_at less_imp_le dist_commute) moreover have "0 \ existence_ivl0 x" by (simp add: \x \ X\) have "\\<^sub>F y in at t within existence_ivl0 x. dist (mvar.flow x 0 id_blinfun y) (mvar.flow x 0 id_blinfun t) < e" using mvar.flow_continuous_on[of x 0 id_blinfun] using \0 < e\ tx by (auto simp add: continuous_on one_exivl dest!: tendstoD) then have "\\<^sub>F (y, s) in at (x, t) within ?S. dist (Dflow x s) (Dflow x t) < e" using \0 < e\ unfolding at_within_open[OF \(x, t) \ ?S\ open_state_space] UNIV_Times_UNIV[symmetric] Dflow_def by (intro eventually_at_Pair_within_TimesI2) (auto simp: at_within_open[OF tx open_existence_ivl]) ultimately have "\\<^sub>F (y, s) in at (x, t) within ?S. dist (Dflow y s) (Dflow x t) < e'" apply eventually_elim proof (safe del: UnE, goal_cases) case (1 y s) have "dist (Dflow y s) (Dflow x t) \ dist (Dflow y s) (Dflow x s) + dist (Dflow x s) (Dflow x t)" by (rule dist_triangle) also have "dist (Dflow x s) (Dflow x t) < e" by (rule 1) also have "dist (Dflow y s) (Dflow x s) < e" unfolding dist_norm norm_minus_commute using 1 by (intro cont_data) also have "e + e \ e'" by (simp add: e_def) finally show "dist (Dflow y s) (Dflow x t) < e'" by arith qed then show "\\<^sub>F ys in at (x, t) within ?S. dist (Dflow (fst ys) (snd ys)) (Dflow (fst (x, t)) (snd (x, t))) < e'" by (simp add: split_beta') qed lemma W_continuous_on_comp[continuous_intros]: assumes h: "continuous_on S h" and g: "continuous_on S g" shows "(\s. s \ S \ h s \ X) \ (\s. s \ S \ g s \ existence_ivl0 (h s)) \ continuous_on S (\s. Dflow (h s) (g s))" using continuous_on_compose[OF continuous_on_Pair[OF h g] continuous_on_subset[OF W_continuous_on]] by auto lemma f_flow_continuous_on: "continuous_on (Sigma X existence_ivl0) (\(x0, t). f (flow0 x0 t))" using flow_continuous_on_state_space by (auto intro!: continuous_on_f flow_in_domain simp: split_beta') lemma flow_has_space_derivative: assumes "t \ existence_ivl0 x0" shows "((\x0. flow0 x0 t) has_derivative Dflow x0 t) (at x0)" by (rule flow_dx_derivative_blinfun[THEN has_derivative_eq_rhs]) (simp_all add: var_eq_mvar assms blinfun.blinfun_apply_inverse Dflow_def vector_Dflow_def mem_existence_ivl_iv_defined[OF assms]) lemma flow_has_flowderiv: assumes "t \ existence_ivl0 x0" shows "((\(x0, t). flow0 x0 t) has_derivative flowderiv x0 t) (at (x0, t) within S)" proof - have Sigma: "(x0, t) \ Sigma X existence_ivl0" using assms by auto from open_state_space assms obtain e' where e': "e' > 0" "ball (x0, t) e' \ Sigma X existence_ivl0" by (force simp: open_contains_ball) define e where "e = e' / sqrt 2" have "0 < e" using e' by (auto simp: e_def) have "ball x0 e \ ball t e \ ball (x0, t) e'" by (auto simp: dist_prod_def real_sqrt_sum_squares_less e_def) also note e'(2) finally have subs: "ball x0 e \ ball t e \ Sigma X existence_ivl0" . have d1: "((\x0. flow0 x0 s) has_derivative blinfun_apply (Dflow y s)) (at y within ball x0 e)" if "y \ ball x0 e" "s \ ball t e" for y s using subs that by (subst at_within_open; force intro!: flow_has_space_derivative) have d2: "(flow0 y has_derivative blinfun_apply (blinfun_scaleR_left (f (flow0 y s)))) (at s within ball t e)" if "y \ ball x0 e" "s \ ball t e" for y s using subs that unfolding has_vector_derivative_eq_has_derivative_blinfun[symmetric] by (subst at_within_open; force intro!: flow_has_vector_derivative) have "((\(x0, t). flow0 x0 t) has_derivative flowderiv x0 t) (at (x0, t) within ball x0 e \ ball t e)" using subs unfolding UNIV_Times_UNIV[symmetric] by (intro has_derivative_partialsI[OF d1 d2, THEN has_derivative_eq_rhs]) (auto intro!: \0 < e\ continuous_intros flow_in_domain continuous_on_imp_continuous_within[where s="Sigma X existence_ivl0"] assms simp: flowderiv_def split_beta' flow0_defined assms mem_ball) then have "((\(x0, t). flow0 x0 t) has_derivative flowderiv x0 t) (at (x0, t) within Sigma X existence_ivl0)" by (auto simp: at_within_open[OF _ open_state_space] at_within_open[OF _ open_Times] assms \0 < e\ mem_existence_ivl_iv_defined[OF assms]) then show ?thesis unfolding at_within_open[OF Sigma open_state_space] by (rule has_derivative_at_withinI) qed lemma flow0_comp_has_derivative: assumes h: "h s \ existence_ivl0 (g s)" assumes [derivative_intros]: "(g has_derivative g') (at s within S)" assumes [derivative_intros]: "(h has_derivative h') (at s within S)" shows "((\x. flow0 (g x) (h x)) has_derivative (\x. blinfun_apply (flowderiv (g s) (h s)) (g' x, h' x))) (at s within S)" by (rule has_derivative_compose[where f="\x. (g x, h x)" and s=S, OF _ flow_has_flowderiv[OF h], simplified]) (auto intro!: derivative_eq_intros) lemma flowderiv_continuous_on: "continuous_on (Sigma X existence_ivl0) (\(x0, t). flowderiv x0 t)" unfolding flowderiv_def split_beta' by (subst blinfun_of_matrix_works[where f="comp12 (Dflow (fst x) (snd x)) (blinfun_scaleR_left (f (flow0 (fst x) (snd x))))" for x, symmetric]) (auto intro!: continuous_intros flow_in_domain) lemma flowderiv_continuous_on_comp[continuous_intros]: assumes "continuous_on S x" assumes "continuous_on S t" assumes "\s. s \ S \ x s \ X" "\s. s \ S \ t s \ existence_ivl0 (x s)" shows "continuous_on S (\xa. flowderiv (x xa) (t xa))" by (rule continuous_on_compose2[OF flowderiv_continuous_on, where f="\s. (x s, t s)", unfolded split_beta' fst_conv snd_conv]) (auto intro!: continuous_intros assms) lemmas [intro] = flow_in_domain lemma vareq_trans: "t0 \ existence_ivl0 x0 \ t \ existence_ivl0 (flow0 x0 t0) \ vareq (flow0 x0 t0) t = vareq x0 (t0 + t)" by (auto simp: vareq_def flow_trans) lemma diff_existence_ivl_trans: "t0 \ existence_ivl0 x0 \ t \ existence_ivl0 x0 \ t - t0 \ existence_ivl0 (flow0 x0 t0)" for t by (metis (no_types, opaque_lifting) add.left_neutral diff_add_eq local.existence_ivl_reverse local.existence_ivl_trans local.flows_reverse) lemma has_vderiv_on_blinfun_compose_right[derivative_intros]: assumes "(g has_vderiv_on g') T" assumes "\x. x \ T \ gd' x = g' x o\<^sub>L d" shows "((\x. g x o\<^sub>L d) has_vderiv_on gd') T" using assms by (auto simp: has_vderiv_on_def has_vector_derivative_def blinfun_ext blinfun.bilinear_simps intro!: derivative_eq_intros ext) lemma has_vderiv_on_blinfun_compose_left[derivative_intros]: assumes "(g has_vderiv_on g') T" assumes "\x. x \ T \ gd' x = d o\<^sub>L g' x" shows "((\x. d o\<^sub>L g x) has_vderiv_on gd') T" using assms by (auto simp: has_vderiv_on_def has_vector_derivative_def blinfun_ext blinfun.bilinear_simps intro!: derivative_eq_intros ext) lemma mvar_flow_shift: assumes "t0 \ existence_ivl0 x0" "t1 \ existence_ivl0 x0" shows "mvar.flow x0 t0 d t1 = Dflow (flow0 x0 t0) (t1 - t0) o\<^sub>L d" proof - have "mvar.flow x0 t0 d t1 = mvar.flow x0 t0 d (t0 + (t1 - t0))" by simp also have "\ = mvar.flow x0 t0 (mvar.flow x0 t0 d t0) t1" by (subst mvar.flow_trans) (auto simp add: assms) also have "\ = Dflow (flow0 x0 t0) (t1 - t0) o\<^sub>L d" apply (rule mvar.flow_unique_on) apply (auto simp add: assms mvar.flow_initial_time_if blinfun_ext Dflow_def intro!: derivative_intros derivative_eq_intros) apply (auto simp: assms has_vderiv_on_open has_vector_derivative_def intro!: derivative_eq_intros blinfun_eqI) apply (subst mvar_existence_ivl_eq_existence_ivl) by (auto simp add: vareq_trans assms diff_existence_ivl_trans) finally show ?thesis . qed lemma Dflow_trans: assumes "h \ existence_ivl0 x0" assumes "i \ existence_ivl0 (flow0 x0 h)" shows "Dflow x0 (h + i) = Dflow (flow0 x0 h) i o\<^sub>L (Dflow x0 h)" proof - have [intro, simp]: "h + i \ existence_ivl0 x0" "i + h \ existence_ivl0 x0" "x0 \ X" using assms by (auto simp add: add.commute existence_ivl_trans) show ?thesis unfolding Dflow_def apply (subst mvar.flow_trans[where s=h and t=i]) subgoal by (auto simp: assms) subgoal by (auto simp: assms) by (subst mvar_flow_shift) (auto simp: assms Dflow_def ) qed lemma Dflow_trans_apply: assumes "h \ existence_ivl0 x0" assumes "i \ existence_ivl0 (flow0 x0 h)" shows "Dflow x0 (h + i) d0 = Dflow (flow0 x0 h) i (Dflow x0 h d0)" proof - have [intro, simp]: "h + i \ existence_ivl0 x0" "i + h \ existence_ivl0 x0" "x0 \ X" using assms by (auto simp add: add.commute existence_ivl_trans) show ?thesis unfolding Dflow_def apply (subst mvar.flow_trans[where s=h and t=i]) subgoal by (auto simp: assms) subgoal by (auto simp: assms) by (subst mvar_flow_shift) (auto simp: assms Dflow_def ) qed end \ \@{thm c1_on_open_euclidean_anchor}\ end