diff --git a/thys/Bernoulli/Bernoulli_Zeta.thy b/thys/Bernoulli/Bernoulli_Zeta.thy --- a/thys/Bernoulli/Bernoulli_Zeta.thy +++ b/thys/Bernoulli/Bernoulli_Zeta.thy @@ -1,502 +1,502 @@ section \Bernoulli numbers and the zeta function at positive integers\ theory Bernoulli_Zeta imports - "HOL-Analysis.Analysis" + "HOL-Complex_Analysis.Complex_Analysis" Bernoulli_FPS begin (* TODO: Move *) lemma joinpaths_cong: "f = f' \ g = g' \ f +++ g = f' +++ g'" by simp lemma linepath_cong: "a = a' \ b = b' \ linepath a b = linepath a' b'" by simp text \ The analytic continuation of the exponential generating function of the Bernoulli numbers is $\frac{z}{e^z - 1}$, which has simple poles at all $2ki\pi$ for $k\in\mathbb{Z}\setminus\{0\}$. We will need the residue at these poles: \ lemma residue_bernoulli: assumes "n \ 0" shows "residue (\z. 1 / (z ^ m * (exp z - 1))) (2 * pi * real_of_int n * \) = 1 / (2 * pi * real_of_int n * \) ^ m" proof - have "residue (\z. (1 / z ^ m) / (exp z - 1)) (2 * pi * real_of_int n * \) = 1 / (2 * pi * real_of_int n * \) ^ m / 1" using exp_integer_2pi[of "real_of_int n"] and assms by (rule_tac residue_simple_pole_deriv[where s="-{0}"]) (auto intro!: holomorphic_intros derivative_eq_intros connected_open_delete_finite simp add: mult_ac connected_punctured_universe) thus ?thesis by (simp add: divide_simps) qed text \ At positive integers greater than 1, the Riemann zeta function is simply the infinite sum $\zeta(n) = \sum_{k=1}^\infty k^{-n}$. For even $n$, this quantity can also be expressed in terms of Bernoulli numbers. To show this, we employ a similar strategy as in the meromorphic asymptotics approach: We apply the Residue Theorem to the exponential generating function of the Bernoulli numbers: \[\sum_{n=0}^\infty \frac{B_n}{n!} z^n = \frac{z}{e^z - 1}\] Recall that this function has poles at $2ki\pi$ for $k\in\mathbb{Z}\setminus\{0\}$. In the meromorphic asymptotics case, we integrated along a circle of radius $3i\pi$ in order to get the dominant singularities $2i\pi$ and $-2i\pi$. Now, however, we will not use a fixed integration path, but we let the integration path become bigger and bigger. Because the integrand decays relatively quickly if $n > 1$, the integral vanishes in the limit and we obtain not just an asymptotic formula, but an exact representation of $B_n$ as an infinite sum. For odd $n$, we have $B_n = 0$, but for even $n$, the residues at $2ki\pi$ and $-2ki\pi$ combine nicely to $2\cdot(-2k\pi)^{-n}$, and after some simplification we get the formula for $B_n$. Another difference to the meromorphic asymptotics is that we now use a rectangle instead of a circle as the integration path. For the asymptotics, only a big-oh bound was needed for the integral over one fixed integration path, and the circular path was very convenient. However, now we need to explicitly bound the integral for a whole sequence of integration paths that grow in size, and bounding $e^z - 1$ for $z$ on a circle is very tedious. On a rectangle, this term can be bounded much more easily. Still, we have to do this separately for all four edges of the rectangle, which will be a bit tedious. \ theorem nat_even_power_sums_complex: assumes n': "n' > 0" shows "(\k. 1 / of_nat (Suc k) ^ (2*n') :: complex) sums of_real ((-1) ^ Suc n' * bernoulli (2*n') * (2 * pi) ^ (2 * n') / (2 * fact (2*n')))" proof - define n where "n = 2 * n'" from n' have n: "n \ 2" "even n" by (auto simp: n_def) define zeta :: complex where "zeta = (\k. 1 / of_nat (Suc k) ^ n)" have "summable (\k. 1 / of_nat (Suc k) ^ n :: complex)" using inverse_power_summable[of n] n by (subst summable_Suc_iff) (simp add: divide_simps) hence "(\k. \i zeta" by (subst (asm) summable_sums_iff) (simp add: sums_def zeta_def) also have "(\k. \ik. \i\{0<..k}. 1 / of_nat i ^ n)" by (intro ext sum.reindex_bij_witness[of _ "\n. n - 1" Suc]) auto finally have zeta_limit: "(\k. \i\{0<..k}. 1 / of_nat i ^ n) \ zeta" . \ \This is the exponential generating function of the Bernoulli numbers.\ define f where "f = (\z::complex. if z = 0 then 1 else z / (exp z - 1))" \ \We will integrate over this function, since its residue at the origin is the $n$-th coefficient of @{term f}. Note that it has singularities at all points $2ik\pi$ for $k\in\mathbb{Z}$.\ define g where "g = (\z::complex. 1 / (z ^ n * (exp z - 1)))" \ \We integrate along a rectangle of width $2m$ and height $2(2m+1)\pi$ with its centre at the origin. The benefit of the rectangular path is that it is easier to bound the value of the exponential appearing in the integrand. The horizontal lines of the rectangle are always right in the middle between two adjacent singularities.\ define \ :: "nat \ real \ complex" where "\ = (\m. rectpath (-real m - real (2*m+1)*pi*\) (real m + real (2*m+1)*pi*\))" \ \This set is a convex open enclosing set the contains our path.\ define A where "A = (\m::nat. box (-(real m+1) - (2*m+2)*pi*\) (real m+1 + (2*m+2)*pi*\))" \ \These are all the singularities in the enclosing inside the path (and also inside @{term A}).\ define S where "S = (\m::nat. (\n. 2 * pi * of_int n * \) ` {-m..m})" \ \Any singularity in @{term A} is of the form $2ki\pi$ where $|k| \leq m$.\ have int_bound: "k \ {-int m..int m}" if "2 * pi * k * \ \ A m" for k m proof - from that have "(-real (Suc m)) * (2 * pi) < real_of_int k * (2 * pi) \ real (Suc m) * (2 * pi) > real_of_int k * (2 * pi)" by (auto simp: A_def in_box_complex_iff algebra_simps) hence "-real (Suc m) < real_of_int k \ real_of_int k < real (Suc m)" by simp also have "-real (Suc m) = real_of_int (-int (Suc m))" by simp also have "real (Suc m) = real_of_int (int (Suc m))" by simp also have "real_of_int (- int (Suc m)) < real_of_int k \ real_of_int k < real_of_int (int (Suc m)) \ k \ {-int m..int m}" by (subst of_int_less_iff) auto finally show "k \ {-int m..int m}" . qed have zeros: "\k\{-int m..int m}. z = 2 * pi * of_int k * \" if "z \ A m" "exp z = 1" for z m proof - from that(2) obtain k where z_eq: "z = 2 * pi * of_int k * \" unfolding exp_eq_1 by (auto simp: complex_eq_iff) with int_bound[of k] and that(1) show ?thesis by auto qed have zeros': "z ^ n * (exp z - 1) \ 0" if "z \ A m - S m" for z m using zeros[of z] that by (auto simp: S_def) \ \The singularities all lie strictly inside the integration path.\ have subset: "S m \ box (-real m - real(2*m+1)*pi*\) (real m + real(2*m+1)*pi*\)" if "m > 0" for m proof (rule, goal_cases) case (1 z) then obtain k :: int where k: "k \ {-int m..int m}" "z = 2 * pi * k * \" unfolding S_def by blast have "2 * pi * -m + -pi < 2 * pi * k + 0" using k by (intro add_le_less_mono mult_left_mono) auto moreover have "2 * pi * k + 0 < 2 * pi * m + pi" using k by (intro add_le_less_mono mult_left_mono) auto ultimately show ?case using k \m > 0\ by (auto simp: A_def in_box_complex_iff algebra_simps) qed from n and zeros' have holo: "g holomorphic_on A m - S m" for m unfolding g_def by (intro holomorphic_intros) auto \ \The integration path lies completely inside $A$ and does not cross any singularities.\ have path_subset: "path_image (\ m) \ A m - S m" if "m > 0" for m proof - have "path_image (\ m) \ cbox (-real m - (2 * m + 1) * pi * \) (real m + (2 * m + 1) * pi * \)" unfolding \_def by (rule path_image_rectpath_subset_cbox) auto also have "\ \ A m" unfolding A_def by (subst subset_box_complex) auto finally have "path_image (\ m) \ A m" . moreover have "path_image (\ m) \ S m = {}" proof safe fix z assume z: "z \ path_image (\ m)" "z \ S m" from this(2) obtain k :: int where k: "z = 2 * pi * k * \" by (auto simp: S_def) hence [simp]: "Re z = 0" by simp from z(1) have "\Im z\ = of_int (2*m+1) * pi" using \m > 0\ by (auto simp: \_def path_image_rectpath) also have "\Im z\ = of_int (2 * \k\) * pi" by (simp add: k abs_mult) finally have "2 * \k\ = 2 * m + 1" by (subst (asm) mult_cancel_right, subst (asm) of_int_eq_iff) simp hence False by presburger thus "z \ {}" .. qed ultimately show "path_image (\ m) \ A m - S m" by blast qed \ \We now obtain a closed form for the Bernoulli numbers using the integral.\ have eq: "(\x\{0<..m}. 1 / of_nat x ^ n) = contour_integral (\ m) g * (2 * pi * \) ^ n / (4 * pi * \) - complex_of_real (bernoulli n / fact n) * (2 * pi * \) ^ n / 2" if m: "m > 0" for m proof - \ \We relate the formal power series of the Bernoulli numbers to the corresponding complex function.\ have "subdegree (fps_exp 1 - 1 :: complex fps) = 1" by (intro subdegreeI) auto hence expansion: "f has_fps_expansion bernoulli_fps" unfolding f_def bernoulli_fps_def by (auto intro!: fps_expansion_intros) \ \We use the Residue Theorem to explicitly compute the integral.\ have "contour_integral (\ m) g = 2 * pi * \ * (\z\S m. winding_number (\ m) z * residue g z)" proof (rule Residue_theorem) have "cbox (-real m - (2 * m + 1) * pi * \) (real m + (2 * m + 1) * pi * \) \ A m" unfolding A_def by (subst subset_box_complex) simp_all thus "\z. z \ A m \ winding_number (\ m) z = 0" unfolding \_def by (intro winding_number_rectpath_outside allI impI) auto qed (insert holo path_subset m, auto simp: \_def A_def S_def intro: convex_connected) \ \Clearly, all the winding numbers are 1\ also have "winding_number (\ m) z = 1" if "z \ S m" for z unfolding \_def using subset[of m] that m by (subst winding_number_rectpath) blast+ hence "(\z\S m. winding_number (\ m) z * residue g z) = (\z\S m. residue g z)" by (intro sum.cong) simp_all also have "\ = (\k=-int m..int m. residue g (2 * pi * of_int k * \))" unfolding S_def by (subst sum.reindex) (auto simp: inj_on_def o_def) also have "{-int m..int m} = insert 0 ({-int m..int m}-{0})" by auto also have "(\k\\. residue g (2 * pi * of_int k * \)) = residue g 0 + (\k\{-int m..m}-{0}. residue g (2 * pi * of_int k * \))" by (subst sum.insert) auto \ \The residue at the origin is just the $n$-th coefficient of $f$.\ also have "residue g 0 = residue (\z. f z / z ^ Suc n) 0" unfolding f_def g_def by (intro residue_cong eventually_mono[OF eventually_at_ball[of 1]]) auto also have "\ = fps_nth bernoulli_fps n" by (rule residue_fps_expansion_over_power_at_0 [OF expansion]) also have "\ = of_real (bernoulli n / fact n)" by simp also have "(\k\{-int m..m}-{0}. residue g (2 * pi * of_int k * \)) = (\k\{-int m..m}-{0}. 1 / of_int k ^ n) / (2 * pi * \) ^ n" proof (subst sum_divide_distrib, intro refl sum.cong, goal_cases) case (1 k) hence *: "residue g (2 * pi * of_int k * \) = 1 / (2 * complex_of_real pi * of_int k * \) ^ n" unfolding g_def by (subst residue_bernoulli) auto thus ?case using 1 by (subst *) (simp add: divide_simps power_mult_distrib) qed also have "(\k\{-int m..m}-{0}. 1 / of_int k ^ n) = (\(a,b)\{0<..m}\{-1,1::int}. 1 / of_int (int a) ^ n :: complex)" using n by (intro sum.reindex_bij_witness[of _ "\k. snd k * int (fst k)" "\k. (nat \k\,sgn k)"]) (auto split: if_splits simp: abs_if) also have "\ = (\x\{0<..m}. 2 / of_nat x ^ n)" using n by (subst sum.Sigma [symmetric]) auto also have "\ = (\x\{0<..m}. 1 / of_nat x ^ n) * 2" by (simp add: sum_distrib_right) finally show ?thesis by (simp add: field_simps) qed \ \The ugly part: We have to prove a bound on the integral by splitting it into four integrals over lines and bounding each part separately.\ have "eventually (\m. norm (contour_integral (\ m) g) \ ((4 + 12 * pi) + 6 * pi / m) / real m ^ (n - 1)) sequentially" using eventually_gt_at_top[of "1::nat"] proof eventually_elim case (elim m) let ?c = "(2*m+1) * pi * \" define I where "I = (\p1 p2. contour_integral (linepath p1 p2) g)" define p1 p2 p3 p4 where "p1 = -real m - ?c" and "p2 = real m - ?c" and "p3 = real m + ?c" and "p4 = -real m + ?c" have eq: "\ m = linepath p1 p2 +++ linepath p2 p3 +++ linepath p3 p4 +++ linepath p4 p1" (is "\ m = ?\'") unfolding \_def rectpath_def Let_def by (intro joinpaths_cong linepath_cong) (simp_all add: p1_def p2_def p3_def p4_def complex_eq_iff) have integrable: "g contour_integrable_on \ m" using elim by (intro contour_integrable_holomorphic_simple[OF holo _ _ path_subset]) (auto simp: \_def A_def S_def intro!: finite_imp_closed) have "norm (contour_integral (\ m) g) = norm (I p1 p2 + I p2 p3 + I p3 p4 + I p4 p1)" unfolding I_def by (insert integrable, unfold eq) (subst contour_integral_join; (force simp: add_ac)?)+ also have "\ \ norm (I p1 p2) + norm (I p2 p3) + norm (I p3 p4) + norm (I p4 p1)" by (intro norm_triangle_mono order.refl) also have "norm (I p1 p2) \ 1 / real m ^ n * norm (p2 - p1)" (is "_ \ ?B1 * _") unfolding I_def proof (intro contour_integral_bound_linepath) fix z assume z: "z \ closed_segment p1 p2" define a where "a = Re z" from z have z: "z = a - (2*m+1) * pi * \" by (subst (asm) closed_segment_same_Im) (auto simp: p1_def p2_def complex_eq_iff a_def) have "real m * 1 \ (2*m+1) * pi" using pi_ge_two by (intro mult_mono) auto also have "(2*m+1) * pi = \Im z\" by (simp add: z) also have "\Im z\ \ norm z" by (rule abs_Im_le_cmod) finally have "norm z \ m" by simp moreover { have "exp z - 1 = -of_real (exp a + 1)" using exp_integer_2pi_plus1[of m] by (simp add: z exp_diff algebra_simps exp_of_real) also have "norm \ \ 1" unfolding norm_minus_cancel norm_of_real by simp finally have "norm (exp z - 1) \ 1" . } ultimately have "norm z ^ n * norm (exp z - 1) \ real m ^ n * 1" by (intro mult_mono power_mono) auto thus "norm (g z) \ 1 / real m ^ n" using elim by (simp add: g_def divide_simps norm_divide norm_mult norm_power mult_less_0_iff) qed (insert integrable, auto simp: eq) also have "norm (p2 - p1) = 2 * m" by (simp add: p2_def p1_def) also have "norm (I p3 p4) \ 1 / real m ^ n * norm (p4 - p3)" (is "_ \ ?B3 * _") unfolding I_def proof (intro contour_integral_bound_linepath) fix z assume z: "z \ closed_segment p3 p4" define a where "a = Re z" from z have z: "z = a + (2*m+1) * pi * \" by (subst (asm) closed_segment_same_Im) (auto simp: p3_def p4_def complex_eq_iff a_def) have "real m * 1 \ (2*m+1) * pi" using pi_ge_two by (intro mult_mono) auto also have "(2*m+1) * pi = \Im z\" by (simp add: z) also have "\Im z\ \ norm z" by (rule abs_Im_le_cmod) finally have "norm z \ m" by simp moreover { have "exp z - 1 = -of_real (exp a + 1)" using exp_integer_2pi_plus1[of m] by (simp add: z exp_add algebra_simps exp_of_real) also have "norm \ \ 1" unfolding norm_minus_cancel norm_of_real by simp finally have "norm (exp z - 1) \ 1" . } ultimately have "norm z ^ n * norm (exp z - 1) \ real m ^ n * 1" by (intro mult_mono power_mono) auto thus "norm (g z) \ 1 / real m ^ n" using elim by (simp add: g_def divide_simps norm_divide norm_mult norm_power mult_less_0_iff) qed (insert integrable, auto simp: eq) also have "norm (p4 - p3) = 2 * m" by (simp add: p4_def p3_def) also have "norm (I p2 p3) \ (1 / real m ^ n) * norm (p3 - p2)" (is "_ \ ?B2 * _") unfolding I_def proof (rule contour_integral_bound_linepath) fix z assume z: "z \ closed_segment p2 p3" define b where "b = Im z" from z have z: "z = m + b * \" by (subst (asm) closed_segment_same_Re) (auto simp: p2_def p3_def algebra_simps complex_eq_iff b_def) from elim have "2 \ 1 + real m" by simp also have "\ \ exp (real m)" by (rule exp_ge_add_one_self) also have "exp (real m) - 1 = norm (exp z) - norm (1::complex)" by (simp add: z) also have "\ \ norm (exp z - 1)" by (rule norm_triangle_ineq2) finally have "norm (exp z - 1) \ 1" by simp moreover have "norm z \ m" using z and abs_Re_le_cmod[of z] by simp ultimately have "norm z ^ n * norm (exp z - 1) \ real m ^ n * 1" using elim by (intro mult_mono power_mono) (auto simp: z) thus "norm (g z) \ 1 / real m ^ n" using n and elim by (simp add: g_def norm_mult norm_divide norm_power divide_simps mult_less_0_iff) qed (insert integrable, auto simp: eq) also have "p3 - p2 = of_real (2*(2*real m+1)*pi) * \" by (simp add: p2_def p3_def) also have "norm \ = 2 * (2 * real m + 1) * pi" unfolding norm_mult norm_of_real by simp also have "norm (I p4 p1) \ (2 / real m ^ n) * norm (p1 - p4)" (is "_ \ ?B4 * _") unfolding I_def proof (rule contour_integral_bound_linepath) fix z assume z: "z \ closed_segment p4 p1" define b where "b = Im z" from z have z: "z = -real m + b * \" by (subst (asm) closed_segment_same_Re) (auto simp: p1_def p4_def algebra_simps b_def complex_eq_iff) from elim have "2 \ 1 + real m" by simp also have "\ \ exp (real m)" by (rule exp_ge_add_one_self) finally have "1 / 2 \ 1 - exp (-real m)" by (subst exp_minus) (simp add: field_simps) also have "1 - exp (-real m) = norm (1::complex) - norm (exp z)" by (simp add: z) also have "\ \ norm (exp z - 1)" by (subst norm_minus_commute, rule norm_triangle_ineq2) finally have "norm (exp z - 1) \ 1 / 2" by simp moreover have "norm z \ m" using z and abs_Re_le_cmod[of z] by simp ultimately have "norm z ^ n * norm (exp z - 1) \ real m ^ n * (1 / 2)" using elim by (intro mult_mono power_mono) (auto simp: z) thus "norm (g z) \ 2 / real m ^ n" using n and elim by (simp add: g_def norm_mult norm_divide norm_power divide_simps mult_less_0_iff) qed (insert integrable, auto simp: eq) also have "p1 - p4 = -of_real (2*(2*real m+1)*pi) * \" by (simp add: p1_def p4_def algebra_simps) also have "norm \ = 2 * (2 * real m + 1) * pi" unfolding norm_mult norm_of_real norm_minus_cancel by simp also have "?B1 * (2*m) + ?B2 * (2*(2*real m+1)*pi) + ?B3 * (2*m) + ?B4 * (2*(2*real m+1)*pi) = (4 * m + 6 * (2 * m + 1) * pi) / real m ^ n" by (simp add: divide_simps) also have "(4 * m + 6 * (2 * m + 1) * pi) = (4 + 12 * pi) * m + 6 * pi" by (simp add: algebra_simps) also have "\ / real m ^ n = ((4 + 12 * pi) + 6 * pi / m) / real m ^ (n - 1)" using n by (cases n) (simp_all add: divide_simps) finally show "cmod (contour_integral (\ m) g) \ \" by simp qed \ \It is clear that this bound goes to 0 since @{prop "n \ 2"}.\ moreover have "(\m. (4 + 12 * pi + 6 * pi / real m) / real m ^ (n - 1)) \ 0" by (rule real_tendsto_divide_at_top tendsto_add tendsto_const filterlim_real_sequentially filterlim_pow_at_top | use n in simp)+ ultimately have *: "(\m. contour_integral (\ m) g) \ 0" by (rule Lim_null_comparison) \ \Since the infinite sum over the residues can expressed using the zeta function, we have now related the Bernoulli numbers at even positive integers to the zeta function.\ have "(\m. contour_integral (\ m) g * (2 * pi * \) ^ n / (4 * pi * \) - of_real (bernoulli n / fact n) * (2 * pi * \) ^ n / 2) \ 0 * (2 * pi * \) ^ n / (4 * pi * \) - of_real (bernoulli n / fact n) * (2 * pi * \) ^ n / 2" using n by (intro tendsto_intros * zeta_limit) auto also have "?this \ (\m. \k\{0<..m}. 1 / of_nat k ^ n) \ - of_real (bernoulli n / fact n) * (2 * pi * \) ^ n / 2" by (intro filterlim_cong eventually_mono [OF eventually_gt_at_top[of "0::nat"]]) (use eq in simp_all) finally have "(\m. \k\{0<..m}. 1 / of_nat k ^ n) \ - of_real (bernoulli n / fact n) * (of_real (2 * pi) * \) ^ n / 2" (is "_ \ ?L") . also have "(\m. \k\{0<..m}. 1 / of_nat k ^ n) = (\m. \k\{..n. n - 1"]) auto also have "\ \ ?L \ (\k. 1 / of_nat (Suc k) ^ n) sums ?L" by (simp add: sums_def) also have "(2 * pi * \) ^ n = (2 * pi) ^ n * (-1) ^ n'" by (simp add: n_def divide_simps power_mult_distrib power_mult power_minus') also have "- of_real (bernoulli n / fact n) * \ / 2 = of_real ((-1) ^ Suc n' * bernoulli (2*n') * (2*pi)^(2*n') / (2 * fact (2*n')))" by (simp add: n_def divide_simps) finally show ?thesis unfolding n_def . qed corollary nat_even_power_sums_real: assumes n': "n' > 0" shows "(\k. 1 / real (Suc k) ^ (2*n')) sums ((-1) ^ Suc n' * bernoulli (2*n') * (2 * pi) ^ (2 * n') / (2 * fact (2*n')))" (is "?f sums ?L") proof - have "(\k. complex_of_real (?f k)) sums complex_of_real ?L" using nat_even_power_sums_complex[OF assms] by simp thus ?thesis by (simp only: sums_of_real_iff) qed text \ We can now also easily determine the signs of Bernoulli numbers: the above formula clearly shows that the signs of $B_{2n}$ alternate as $n$ increases, and we already know that $B_{2n+1} = 0$ for any positive $n$. A lot of other facts about the signs of Bernoulli numbers follow. \ corollary sgn_bernoulli_even: assumes "n > 0" shows "sgn (bernoulli (2 * n)) = (-1) ^ Suc n" proof - have *: "(\k. 1 / real (Suc k) ^ (2 * n)) sums ((- 1) ^ Suc n * bernoulli (2 * n) * (2 * pi) ^ (2 * n) / (2 * fact (2 * n)))" using assms by (rule nat_even_power_sums_real) from * have "0 < (\k. 1 / real (Suc k) ^ (2*n))" by (intro suminf_pos) (auto simp: sums_iff) hence "sgn (\k. 1 / real (Suc k) ^ (2*n)) = 1" by simp also have "(\k. 1 / real (Suc k) ^ (2*n)) = (- 1) ^ Suc n * bernoulli (2 * n) * (2 * pi) ^ (2 * n) / (2 * fact (2 * n))" using * by (simp add: sums_iff) also have "sgn \ = (-1) ^ Suc n * sgn (bernoulli (2 * n))" by (simp add: sgn_mult) finally show ?thesis by (simp add: minus_one_power_iff split: if_splits) qed corollary bernoulli_even_nonzero: "even n \ bernoulli n \ 0" using sgn_bernoulli_even[of "n div 2"] by (cases "n = 0") (auto elim!: evenE) corollary sgn_bernoulli: "sgn (bernoulli n) = (if n = 0 then 1 else if n = 1 then -1 else if odd n then 0 else (-1) ^ Suc (n div 2))" using sgn_bernoulli_even [of "n div 2"] by (auto simp: bernoulli_odd_eq_0) corollary bernoulli_zero_iff: "bernoulli n = 0 \ odd n \ n \ 1" by (auto simp: bernoulli_even_nonzero bernoulli_odd_eq_0) corollary bernoulli'_zero_iff: "(bernoulli' n = 0) \ (n \ 1 \ odd n)" by (auto simp: bernoulli'_def bernoulli_zero_iff) corollary bernoulli_pos_iff: "bernoulli n > 0 \ n = 0 \ n mod 4 = 2" proof - have "bernoulli n > 0 \ sgn (bernoulli n) = 1" by (simp add: sgn_if) also have "\ \ n = 0 \ even n \ odd (n div 2)" by (subst sgn_bernoulli) auto also have "even n \ odd (n div 2) \ n mod 4 = 2" by presburger finally show ?thesis . qed corollary bernoulli_neg_iff: "bernoulli n < 0 \ n = 1 \ n > 0 \ 4 dvd n" proof - have "bernoulli n < 0 \ sgn (bernoulli n) = -1" by (simp add: sgn_if) also have "\ \ n = 1 \ n > 0 \ even n \ even (n div 2)" by (subst sgn_bernoulli) (auto simp: minus_one_power_iff) also have "even n \ even (n div 2) \ 4 dvd n" by presburger finally show ?thesis . qed text \ We also get the solution of the Basel problem (the sum over all squares of positive integers) and any `Basel-like' problem with even exponent. The case of odd exponents is much more complicated and no similarly nice closed form is known for these. \ corollary nat_squares_sums: "(\n. 1 / (n+1) ^ 2) sums (pi ^ 2 / 6)" using nat_even_power_sums_real[of 1] by (simp add: fact_numeral) corollary nat_power4_sums: "(\n. 1 / (n+1) ^ 4) sums (pi ^ 4 / 90)" using nat_even_power_sums_real[of 2] by (simp add: fact_numeral) corollary nat_power6_sums: "(\n. 1 / (n+1) ^ 6) sums (pi ^ 6 / 945)" using nat_even_power_sums_real[of 3] by (simp add: fact_numeral) corollary nat_power8_sums: "(\n. 1 / (n+1) ^ 8) sums (pi ^ 8 / 9450)" using nat_even_power_sums_real[of 4] by (simp add: fact_numeral) end diff --git a/thys/Bernoulli/ROOT b/thys/Bernoulli/ROOT --- a/thys/Bernoulli/ROOT +++ b/thys/Bernoulli/ROOT @@ -1,16 +1,16 @@ chapter AFP -session Bernoulli (AFP) = "HOL-Analysis" + +session Bernoulli (AFP) = "HOL-Complex_Analysis" + options [timeout = 600] sessions "HOL-Analysis" "HOL-Number_Theory" "HOL-Computational_Algebra" theories Bernoulli Periodic_Bernpoly Bernoulli_FPS Bernoulli_Zeta document_files "root.tex" "root.bib" diff --git a/thys/Cartan_FP/Cartan.thy b/thys/Cartan_FP/Cartan.thy --- a/thys/Cartan_FP/Cartan.thy +++ b/thys/Cartan_FP/Cartan.thy @@ -1,534 +1,534 @@ theory Cartan -imports "HOL-Analysis.Analysis" +imports "HOL-Complex_Analysis.Complex_Analysis" begin section\First Cartan Theorem\ text\Ported from HOL Light. See Gianni Ciolli, Graziano Gentili, Marco Maggesi. A Certified Proof of the Cartan Fixed Point Theorems. J Automated Reasoning (2011) 47:319--336 DOI 10.1007/s10817-010-9198-6\ lemma deriv_left_inverse: assumes "f holomorphic_on S" and "g holomorphic_on T" and "open S" and "open T" and "f ` S \ T" and [simp]: "\z. z \ S \ g (f z) = z" and "w \ S" shows "deriv f w * deriv g (f w) = 1" proof - have "deriv f w * deriv g (f w) = deriv g (f w) * deriv f w" by (simp add: algebra_simps) also have "... = deriv (g o f) w" using assms - by (metis analytic_on_imp_differentiable_at analytic_on_open complex_derivative_chain image_subset_iff) + by (metis analytic_on_imp_differentiable_at analytic_on_open deriv_chain image_subset_iff) also have "... = deriv id w" apply (rule complex_derivative_transform_within_open [where s=S]) apply (rule assms holomorphic_on_compose_gen holomorphic_intros)+ apply simp done also have "... = 1" using higher_deriv_id [of 1] by simp finally show ?thesis . qed lemma Cauchy_higher_deriv_bound: assumes holf: "f holomorphic_on (ball z r)" and contf: "continuous_on (cball z r) f" and "0 < r" and "0 < n" and fin : "\w. w \ ball z r \ f w \ ball y B0" shows "norm ((deriv ^^ n) f z) \ (fact n) * B0 / r^n" proof - have "0 < B0" using \0 < r\ fin [of z] by (metis ball_eq_empty ex_in_conv fin not_less) have le_B0: "\w. cmod (w - z) \ r \ cmod (f w - y) \ B0" apply (rule continuous_on_closure_norm_le [of "ball z r" "\w. f w - y"]) apply (auto simp: \0 < r\ dist_norm norm_minus_commute) apply (rule continuous_intros contf)+ using fin apply (simp add: dist_commute dist_norm less_eq_real_def) done have "(deriv ^^ n) f z = (deriv ^^ n) (\w. f w) z - (deriv ^^ n) (\w. y) z" using \0 < n\ by simp also have "... = (deriv ^^ n) (\w. f w - y) z" by (rule higher_deriv_diff [OF holf, symmetric]) (auto simp: \0 < r\ holomorphic_on_const) finally have "(deriv ^^ n) f z = (deriv ^^ n) (\w. f w - y) z" . have contf': "continuous_on (cball z r) (\u. f u - y)" by (rule contf continuous_intros)+ have holf': "(\u. (f u - y)) holomorphic_on (ball z r)" by (simp add: holf holomorphic_on_diff holomorphic_on_const) define a where "a = (2 * pi)/(fact n)" have "0 < a" by (simp add: a_def) have "B0/r^(Suc n)*2 * pi * r = a*((fact n)*B0/r^n)" using \0 < r\ by (simp add: a_def divide_simps) have der_dif: "(deriv ^^ n) (\w. f w - y) z = (deriv ^^ n) f z" using \0 < r\ \0 < n\ by (auto simp: higher_deriv_diff [OF holf holomorphic_on_const]) have "norm ((2 * of_real pi * \)/(fact n) * (deriv ^^ n) (\w. f w - y) z) \ (B0/r^(Suc n)) * (2 * pi * r)" apply (rule has_contour_integral_bound_circlepath [of "(\u. (f u - y)/(u - z)^(Suc n))" _ z]) using Cauchy_has_contour_integral_higher_derivative_circlepath [OF contf' holf'] using \0 < B0\ \0 < r\ apply (auto simp: norm_divide norm_mult norm_power divide_simps le_B0) done then show ?thesis using \0 < r\ by (auto simp: norm_divide norm_mult norm_power field_simps der_dif le_B0) qed lemma higher_deriv_comp_lemma: assumes s: "open s" and holf: "f holomorphic_on s" and "z \ s" and t: "open t" and holg: "g holomorphic_on t" and fst: "f ` s \ t" and n: "i \ n" and dfz: "deriv f z = 1" and zero: "\i. \1 < i; i \ n\ \ (deriv ^^ i) f z = 0" shows "(deriv ^^ i) (g o f) z = (deriv ^^ i) g (f z)" using n holg proof (induction i arbitrary: g) case 0 then show ?case by simp next case (Suc i) have "g \ f holomorphic_on s" using "Suc.prems" holf using fst by (simp add: holomorphic_on_compose_gen image_subset_iff) then have 1: "deriv (g \ f) holomorphic_on s" by (simp add: holomorphic_deriv s) have dg: "deriv g holomorphic_on t" using Suc.prems by (simp add: Suc.prems(2) holomorphic_deriv t) then have "deriv g holomorphic_on f ` s" using fst by (simp add: holomorphic_on_subset image_subset_iff) then have dgf: "(deriv g o f) holomorphic_on s" by (simp add: holf holomorphic_on_compose) then have 2: "(\w. (deriv g o f) w * deriv f w) holomorphic_on s" by (blast intro: holomorphic_intros holomorphic_on_compose holf s) have "(deriv ^^ i) (deriv (g o f)) z = (deriv ^^ i) (\w. deriv g (f w) * deriv f w) z" apply (rule higher_deriv_transform_within_open [OF 1 2 [unfolded o_def] s \z \ s\]) - apply (rule complex_derivative_chain) + apply (rule deriv_chain) using holf Suc.prems fst apply (auto simp: holomorphic_on_imp_differentiable_at s t) done also have "... = (\j=0..i. of_nat(i choose j) * (deriv ^^ j) (\w. deriv g (f w)) z * (deriv ^^ (i - j)) (deriv f) z)" apply (rule higher_deriv_mult [OF dgf [unfolded o_def] _ s \z \ s\]) by (simp add: holf holomorphic_deriv s) also have "... = (\j=i..i. of_nat(i choose j) * (deriv ^^ j) (\w. deriv g (f w)) z * (deriv ^^ Suc (i - j)) f z)" proof - have *: "(deriv ^^ j) (\w. deriv g (f w)) z = 0" if "j < i" and nz: "(deriv ^^ (i - j)) (deriv f) z \ 0" for j proof - have "1 < Suc (i - j)" "Suc (i - j) \ n" using \j < i\ \Suc i \ n\ by auto then show ?thesis by (metis comp_def funpow.simps(2) funpow_swap1 zero nz) qed then show ?thesis apply (simp only: funpow_Suc_right o_def) apply (rule comm_monoid_add_class.sum.mono_neutral_right, auto) done qed also have "... = (deriv ^^ i) (deriv g) (f z)" using Suc.IH [OF _ dg] Suc.prems by (simp add: dfz) finally show ?case by (simp only: funpow_Suc_right o_def) qed lemma higher_deriv_comp_iter_lemma: assumes s: "open s" and holf: "f holomorphic_on s" and fss: "f ` s \ s" and "z \ s" and [simp]: "f z = z" and n: "i \ n" and dfz: "deriv f z = 1" and zero: "\i. \1 < i; i \ n\ \ (deriv ^^ i) f z = 0" shows "(deriv ^^ i) (f^^m) z = (deriv ^^ i) f z" proof - have holfm: "(f^^m) holomorphic_on s" for m apply (induction m, simp add: holomorphic_on_ident) apply (simp only: funpow_Suc_right holomorphic_on_compose_gen [OF holf _ fss]) done show ?thesis using n proof (induction m) case 0 with dfz show ?case by (auto simp: zero) next case (Suc m) have "(deriv ^^ i) (f ^^ m \ f) z = (deriv ^^ i) (f ^^ m) (f z)" using Suc.prems holfm \z \ s\ dfz fss higher_deriv_comp_lemma holf s zero by blast also have "... = (deriv ^^ i) f z" by (simp add: Suc) finally show ?case by (simp only: funpow_Suc_right) qed qed lemma higher_deriv_iter_top_lemma: assumes s: "open s" and holf: "f holomorphic_on s" and fss: "f ` s \ s" and "z \ s" and [simp]: "f z = z" and dfz [simp]: "deriv f z = 1" and n: "1 < n" "\i. \1 < i; i < n\ \ (deriv ^^ i) f z = 0" shows "(deriv ^^ n) (f ^^ m) z = m * (deriv ^^ n) f z" using n proof (induction n arbitrary: m) case 0 then show ?case by simp next case (Suc n) have [simp]: "(f^^m) z = z" for m by (induction m) auto have fms_sb: "(f^^m) ` s \ s" for m apply (induction m) using fss apply force+ done have holfm: "(f^^m) holomorphic_on s" for m apply (induction m, simp add: holomorphic_on_ident) apply (simp only: funpow_Suc_right holomorphic_on_compose_gen [OF holf _ fss]) done then have holdfm: "deriv (f ^^ m) holomorphic_on s" for m by (simp add: holomorphic_deriv s) have holdffm: "(\z. deriv f ((f ^^ m) z)) holomorphic_on s" for m apply (rule holomorphic_on_compose_gen [where g="deriv f" and t=s, unfolded o_def]) using s \z \ s\ holfm holf fms_sb by (auto intro: holomorphic_intros) have f_cd_w: "\w. w \ s \ f field_differentiable at w" using holf holomorphic_on_imp_differentiable_at s by blast have f_cd_mw: "\m w. w \ s \ (f^^m) field_differentiable at w" using holfm holomorphic_on_imp_differentiable_at s by auto have der_fm [simp]: "deriv (f ^^ m) z = 1" for m apply (induction m, simp add: deriv_ident) apply (subst funpow_Suc_right) - apply (subst complex_derivative_chain) + apply (subst deriv_chain) using \z \ s\ holfm holomorphic_on_imp_differentiable_at s f_cd_w apply auto done note Suc(3) [simp] note n_Suc = Suc show ?case proof (induction m) case 0 with n_Suc show ?case by (metis Zero_not_Suc funpow_simps_right(1) higher_deriv_id lambda_zero nat_neq_iff of_nat_0) next case (Suc m) have deriv_nffm: "(deriv ^^ n) (deriv f o (f ^^ m)) z = (deriv ^^ n) (deriv f) ((f ^^ m) z)" apply (rule higher_deriv_comp_lemma [OF s holfm \z \ s\ s _ fms_sb order_refl]) using \z \ s\ fss higher_deriv_comp_iter_lemma holf holf holomorphic_deriv s apply auto done have "deriv (f ^^ m \ f) holomorphic_on s" by (metis funpow_Suc_right holdfm) moreover have "(\w. deriv f ((f ^^ m) w) * deriv (f ^^ m) w) holomorphic_on s" by (rule holomorphic_on_mult [OF holdffm holdfm]) ultimately have "(deriv ^^ n) (deriv (f ^^ m \ f)) z = (deriv ^^ n) (\w. deriv f ((f ^^ m) w) * deriv (f ^^ m) w) z" apply (rule higher_deriv_transform_within_open [OF _ _ s \z \ s\]) - by (metis comp_funpow complex_derivative_chain f_cd_mw f_cd_w fms_sb funpow_swap1 image_subset_iff o_id) + by (metis comp_funpow deriv_chain f_cd_mw f_cd_w fms_sb funpow_swap1 image_subset_iff o_id) also have "... = (\i=0..n. of_nat(n choose i) * (deriv ^^ i) (\w. deriv f ((f ^^ m) w)) z * (deriv ^^ (n - i)) (deriv (f ^^ m)) z)" by (rule higher_deriv_mult [OF holdffm holdfm s \z \ s\]) also have "... = (\i \ {0,n}. of_nat(n choose i) * (deriv ^^ i) (\w. deriv f ((f ^^ m) w)) z * (deriv ^^ (n - i)) (deriv (f ^^ m)) z)" proof - have *: "(deriv ^^ i) (\w. deriv f ((f ^^ m) w)) z = 0" if "i \ n" "0 < i" "i \ n" and nz: "(deriv ^^ (n - i)) (deriv (f ^^ m)) z \ 0" for i proof - have less: "1 < Suc (n-i)" and le: "Suc (n-i) \ n" using that by auto have "(deriv ^^ (Suc (n - i))) (f ^^ m) z = (deriv ^^(Suc (n - i))) f z" apply (rule higher_deriv_comp_iter_lemma [OF s holf fss \z \ s\ \f z = z\ le dfz]) by simp also have "... = 0" using n_Suc(3) less le le_imp_less_Suc by blast finally have "(deriv ^^ (Suc (n - i))) (f ^^ m) z = 0" . then show ?thesis by (simp add: funpow_swap1 nz) qed show ?thesis by (rule comm_monoid_add_class.sum.mono_neutral_right) (auto simp: *) qed also have "... = of_nat (Suc m) * (deriv ^^ n) (deriv f) z" apply (subst Groups_Big.comm_monoid_add_class.sum.insert) apply (simp_all add: deriv_nffm [unfolded o_def] of_nat_Suc [of 0] del: of_nat_Suc) using n_Suc(2) Suc apply (auto simp del: funpow.simps simp: algebra_simps funpow_simps_right) done finally have "(deriv ^^ n) (deriv (f ^^ m \ f)) z = of_nat (Suc m) * (deriv ^^ n) (deriv f) z" . then show ?case apply (simp only: funpow_Suc_right) apply (simp add: o_def del: of_nat_Suc) done qed qed text\Should be proved for n-dimensional vectors of complex numbers\ theorem first_Cartan_dim_1: assumes holf: "f holomorphic_on s" and "open s" "connected s" "bounded s" and fss: "f ` s \ s" and "z \ s" and [simp]: "f z = z" and dfz [simp]: "deriv f z = 1" and "w \ s" shows "f w = w" proof - obtain c where "0 < c" and c: "s \ ball z c" using \bounded s\ bounded_subset_ballD by blast obtain r where "0 < r" and r: "cball z r \ s" using \z \ s\ open_contains_cball \open s\ by blast then have bzr: "ball z r \ s" using ball_subset_cball by blast have fms_sb: "(f^^m) ` s \ s" for m apply (induction m) using fss apply force+ done have holfm: "(f^^m) holomorphic_on s" for m apply (induction m, simp add: holomorphic_on_ident) apply (simp only: funpow_Suc_right holomorphic_on_compose_gen [OF holf _ fss]) done have *: "(deriv ^^ n) f z = (deriv ^^ n) id z" for n proof - consider "n = 0" | "n = 1" | "1 < n" by arith then show ?thesis proof cases assume "n = 0" then show ?thesis by force next assume "n = 1" then show ?thesis by force next assume n1: "n > 1" then have "(deriv ^^ n) f z = 0" proof (induction n rule: less_induct) case (less n) have le: "real m * cmod ((deriv ^^ n) f z) \ fact n * c / r ^ n" if "m\0" for m proof - have holfm': "(f ^^ m) holomorphic_on ball z r" using holfm bzr holomorphic_on_subset by blast then have contfm': "continuous_on (cball z r) (f ^^ m)" using \cball z r \ s\ holfm holomorphic_on_imp_continuous_on holomorphic_on_subset by blast have "real m * cmod ((deriv ^^ n) f z) = cmod (real m * (deriv ^^ n) f z)" by (simp add: norm_mult) also have "... = cmod ((deriv ^^ n) (f ^^ m) z)" apply (subst higher_deriv_iter_top_lemma [OF \open s\ holf fss \z \ s\ \f z = z\ dfz]) using less apply auto done also have "... \ fact n * c / r ^ n" apply (rule Cauchy_higher_deriv_bound [OF holfm' contfm' \0 < r\, where y=z]) using less.prems apply linarith using fms_sb c r ball_subset_cball apply blast done finally show ?thesis . qed have "cmod ((deriv ^^ n) f z) = 0" apply (rule real_archimedian_rdiv_eq_0 [where c = "(fact n) * c / r ^ n"]) apply simp using \0 < r\ \0 < c\ apply (simp add: divide_simps) apply (blast intro: le) done then show ?case by simp qed with n1 show ?thesis by simp qed qed have "f w = id w" by (rule holomorphic_fun_eq_on_connected [OF holf holomorphic_on_id \open s\ \connected s\ * \z \ s\ \w \ s\]) also have "... = w" by simp finally show ?thesis . qed text\Second Cartan Theorem.\ lemma Cartan_is_linear: assumes holf: "f holomorphic_on s" and "open s" and "connected s" and "0 \ s" and ins: "\u z. \norm u = 1; z \ s\ \ u * z \ s" and feq: "\u z. \norm u = 1; z \ s\ \ f (u * z) = u * f z" shows "\c. \z \ s. f z = c * z" proof - have [simp]: "f 0 = 0" using feq [of "-1" 0] assms by simp have uneq: "u^n * (deriv ^^ n) f (u * z) = u * (deriv ^^ n) f z" if "norm u = 1" "z \ s" for n u z proof - have holfuw: "(\w. f (u * w)) holomorphic_on s" apply (rule holomorphic_on_compose_gen [OF _ holf, unfolded o_def]) using that ins by (auto simp: holomorphic_on_linear) have hol_d_fuw: "(deriv ^^ n) (\w. u * f w) holomorphic_on s" for n by (rule holomorphic_higher_deriv holomorphic_intros holf assms)+ have *: "(deriv ^^ n) (\w. u * f w) z = u * (deriv ^^ n) f z" if "z \ s" for z using that proof (induction n arbitrary: z) case 0 then show ?case by simp next case (Suc n) have "deriv ((deriv ^^ n) (\w. u * f w)) z = deriv (\w. u * (deriv ^^ n) f w) z" apply (rule complex_derivative_transform_within_open [OF hol_d_fuw]) apply (auto intro!: holomorphic_higher_deriv holomorphic_intros assms Suc) done also have "... = u * deriv ((deriv ^^ n) f) z" apply (rule deriv_cmult) using Suc \open s\ holf holomorphic_higher_deriv holomorphic_on_imp_differentiable_at by blast finally show ?case by simp qed have "(deriv ^^ n) (\w. f (u * w)) z = u ^ n * (deriv ^^ n) f (u * z)" apply (rule higher_deriv_compose_linear [OF holf \open s\ \open s\]) apply (simp add: that) apply (simp add: ins that) done moreover have "(deriv ^^ n) (\w. f (u * w)) z = u * (deriv ^^ n) f z" apply (subst higher_deriv_transform_within_open [OF holfuw, of "\w. u * f w"]) apply (rule holomorphic_intros holf assms that)+ apply blast using * \z \ s\ apply blast done ultimately show ?thesis by metis qed have dnf0: "(deriv ^^ n) f 0 = 0" if len: "2 \ n" for n proof - have **: "z = 0" if "\u::complex. norm u = 1 \ u ^ n * z = u * z" for z proof - have "\u::complex. norm u = 1 \ u ^ n \ u" using complex_not_root_unity [of "n-1"] len apply (simp add: algebra_simps le_diff_conv2, clarify) apply (rule_tac x=u in exI) apply (subst (asm) power_diff) apply auto done with that show ?thesis by auto qed show ?thesis apply (rule **) using uneq [OF _ \0 \ s\] by force qed show ?thesis apply (rule_tac x = "deriv f 0" in exI, clarify) apply (rule holomorphic_fun_eq_on_connected [OF holf _ \open s\ \connected s\ _ \0 \ s\]) using dnf0 apply (auto simp: holomorphic_on_linear) done qed text\Should be proved for n-dimensional vectors of complex numbers\ theorem second_Cartan_dim_1: assumes holf: "f holomorphic_on ball 0 r" and holg: "g holomorphic_on ball 0 r" and [simp]: "f 0 = 0" and [simp]: "g 0 = 0" and ballf: "\z. z \ ball 0 r \ f z \ ball 0 r" and ballg: "\z. z \ ball 0 r \ g z \ ball 0 r" and fg: "\z. z \ ball 0 r \ f (g z) = z" and gf: "\z. z \ ball 0 r \ g (f z) = z" and "0 < r" shows "\t. \z \ ball 0 r. g z = exp(\ * of_real t) * z" proof - have c_le_1: "c \ 1" if "0 \ c" "\x. 0 \ x \ x < r \ c * x < r" for c proof - have rst: "\r s t::real. 0 = r \ s/r < t \ r < 0 \ \ s < r * t" by (metis (no_types) mult_less_cancel_left_disj nonzero_mult_div_cancel_left times_divide_eq_right) { assume "\ r < c \ c * (c * (c * (c * r))) < 1" then have "1 \ c \ (\r. \ 1 < r \ \ r < c)" using \0 \ c\ by (metis (full_types) less_eq_real_def mult.right_neutral mult_left_mono not_less) then have "\ 1 < c \ \ 1 \ c" by linarith } moreover { have "\ 0 \ r / c \ \ 1 \ c" using \0 < r\ by force then have "1 < c \ \ 1 \ c" using rst \0 < r\ that by (metis div_by_1 frac_less2 less_le_trans mult.commute not_le order_refl pos_divide_le_eq zero_less_one) } ultimately show ?thesis by (metis (no_types) linear not_less) qed have ugeq: "u * g z = g (u * z)" if nou: "norm u = 1" and z: "z \ ball 0 r" for u z proof - have [simp]: "u \ 0" using that by auto have hol1: "(\a. f (u * g a) / u) holomorphic_on ball 0 r" apply (rule holomorphic_intros) apply (rule holomorphic_on_compose_gen [OF _ holf, unfolded o_def]) apply (rule holomorphic_intros holg)+ using nou ballg apply (auto simp: dist_norm norm_mult holomorphic_on_const) done have cdf: "f field_differentiable at 0" using \0 < r\ holf holomorphic_on_imp_differentiable_at by auto have cdg: "g field_differentiable at 0" using \0 < r\ holg holomorphic_on_imp_differentiable_at by auto have cd_fug: "(\a. f (u * g a)) field_differentiable at 0" apply (rule field_differentiable_compose [where g=f and f = "\a. (u * g a)", unfolded o_def]) apply (rule derivative_intros)+ using cdf cdg apply auto done have "deriv g 0 = deriv g (f 0)" by simp then have "deriv f 0 * deriv g 0 = 1" by (metis open_ball \0 < r\ ballf centre_in_ball deriv_left_inverse gf holf holg image_subsetI) then have equ: "deriv f 0 * deriv (\a. u * g a) 0 = u" by (simp add: cdg deriv_cmult) have der1: "deriv (\a. f (u * g a) / u) 0 = 1" apply (simp add: field_class.field_divide_inverse deriv_cmult_right [OF cd_fug]) - apply (subst complex_derivative_chain [where g=f and f = "\a. (u * g a)", unfolded o_def]) + apply (subst deriv_chain [where g=f and f = "\a. (u * g a)", unfolded o_def]) apply (rule derivative_intros cdf cdg | simp add: equ)+ done have fugeq: "\w. w \ ball 0 r \ f (u * g w) / u = w" apply (rule first_Cartan_dim_1 [OF hol1, where z=0]) apply (simp_all add: \0 < r\) apply (auto simp: der1) using nou ballf ballg apply (simp add: dist_norm norm_mult norm_divide) done have "f(u * g z) = u * z" by (metis \u \ 0\ fugeq nonzero_mult_div_cancel_left z times_divide_eq_right) also have "... = f (g (u * z))" by (metis (no_types, lifting) fg mem_ball_0 mult_cancel_right2 norm_mult nou z) finally have "f(u * g z) = f (g (u * z))" . then have "g (f (u * g z)) = g (f (g (u * z)))" by simp then show ?thesis apply (subst (asm) gf) apply (simp add: dist_norm norm_mult nou) using ballg mem_ball_0 z apply blast apply (subst (asm) gf) apply (simp add: dist_norm norm_mult nou) apply (metis ballg mem_ball_0 mult.left_neutral norm_mult nou z, simp) done qed obtain c where c: "\z. z \ ball 0 r \ g z = c * z" apply (rule exE [OF Cartan_is_linear [OF holg]]) apply (simp_all add: \0 < r\ ugeq) apply (auto simp: dist_norm norm_mult) done have gr2: "g (f (r/2)) = c * f(r/2)" apply (rule c) using \0 < r\ ballf mem_ball_0 by force then have "norm c > 0" using \0 < r\ by simp (metis \f 0 = 0\ c dist_commute fg mem_ball mult_zero_left perfect_choose_dist) then have [simp]: "c \ 0" by auto have xless: "x < r * cmod c" if "0 \ x" "x < r" for x proof - have "x = norm (g (f (of_real x)))" proof - have "r > cmod (of_real x)" by (simp add: that) then have "complex_of_real x \ ball 0 r" using mem_ball_0 by blast then show ?thesis using gf \0 \ x\ by force qed then show ?thesis apply (rule ssubst) apply (subst c) apply (rule ballf) using ballf [of x] that apply (auto simp: norm_mult dist_0_norm) done qed have 11: "1 / norm c \ 1" apply (rule c_le_1) using xless apply (auto simp: divide_simps) done have "\0 \ x; x < r\ \ cmod c * x < r" for x using c [of x] ballg [of x] by (auto simp: norm_mult dist_0_norm) then have "norm c \ 1" by (force intro: c_le_1) moreover have "1 \ norm c" using 11 by simp ultimately have "norm c = 1" by (rule antisym) with complex_norm_eq_1_exp c show ?thesis by metis qed end diff --git a/thys/Cartan_FP/ROOT b/thys/Cartan_FP/ROOT --- a/thys/Cartan_FP/ROOT +++ b/thys/Cartan_FP/ROOT @@ -1,10 +1,10 @@ chapter AFP -session Cartan_FP (AFP) = "HOL-Analysis" + +session Cartan_FP (AFP) = "HOL-Complex_Analysis" + description "The Cartan Fixed Point Theorems" options [timeout = 300] theories Cartan document_files "root.bib" "root.tex" diff --git a/thys/Count_Complex_Roots/ROOT b/thys/Count_Complex_Roots/ROOT --- a/thys/Count_Complex_Roots/ROOT +++ b/thys/Count_Complex_Roots/ROOT @@ -1,13 +1,13 @@ chapter AFP -session "Count_Complex_Roots" (AFP) = "HOL-Analysis" + +session "Count_Complex_Roots" (AFP) = "HOL-Complex_Analysis" + options [timeout = 600] sessions Sturm_Tarski Winding_Number_Eval theories Count_Complex_Roots Count_Complex_Roots_Examples document_files "root.tex" "root.bib" diff --git a/thys/Dirichlet_Series/Dirichlet_Series_Analysis.thy b/thys/Dirichlet_Series/Dirichlet_Series_Analysis.thy --- a/thys/Dirichlet_Series/Dirichlet_Series_Analysis.thy +++ b/thys/Dirichlet_Series/Dirichlet_Series_Analysis.thy @@ -1,3087 +1,3087 @@ (* File: Dirichlet_Series_Analysis.thy Author: Manuel Eberl, TU München *) section \Analytic properties of Dirichlet series\ theory Dirichlet_Series_Analysis imports - "HOL-Analysis.Analysis" + "HOL-Complex_Analysis.Complex_Analysis" "HOL-Library.Going_To_Filter" "HOL-Real_Asymp.Real_Asymp" Dirichlet_Series Moebius_Mu Partial_Summation Euler_Products begin (* TODO Move *) lemma frequently_eventually_frequently: "frequently P F \ eventually Q F \ frequently (\x. P x \ Q x) F" by (erule frequently_rev_mp, erule eventually_mono) auto text \ The following illustrates a concept we will need later on: A property holds for \f\ going to \F\ if we can find e.\,g.\ a sequence that tends to \F\ and whose elements eventually satisfy \P\. \ lemma frequently_going_toI: assumes "filterlim (\n. f (g n)) F G" assumes "eventually (\n. P (g n)) G" assumes "eventually (\n. g n \ A) G" assumes "G \ bot" shows "frequently P (f going_to F within A)" unfolding frequently_def proof assume "eventually (\x. \P x) (f going_to F within A)" hence "eventually (\x. \P x) (inf (filtercomap f F) (principal A))" by (simp add: going_to_within_def) moreover have "filterlim (\n. g n) (inf (filtercomap f F) (principal A)) G" using assms unfolding filterlim_inf filterlim_principal by (auto simp add: filterlim_iff_le_filtercomap filtercomap_filtercomap) ultimately have "eventually (\n. \P (g n)) G" by (rule eventually_compose_filterlim) with assms(2) have "eventually (\_. False) G" by eventually_elim auto with assms(4) show False by simp qed lemma frequently_filtercomapI: assumes "filterlim (\n. f (g n)) F G" assumes "eventually (\n. P (g n)) G" assumes "G \ bot" shows "frequently P (filtercomap f F)" using frequently_going_toI[of f g F G P UNIV] assms by (simp add: going_to_def) lemma frequently_going_to_at_topE: fixes f :: "'a \ real" assumes "frequently P (f going_to at_top)" obtains g where "\n. P (g n)" and "filterlim (\n. f (g n)) at_top sequentially" proof - from assms have "\k. \x. f x \ real k \ P x" by (auto simp: frequently_def eventually_going_to_at_top_linorder) hence "\g. \k. f (g k) \ real k \ P (g k)" by metis then obtain g where g: "\k. f (g k) \ real k" "\k. P (g k)" by blast have "filterlim (\n. f (g n)) at_top sequentially" by (rule filterlim_at_top_mono[OF filterlim_real_sequentially]) (use g in auto) from g(2) and this show ?thesis using that[of g] by blast qed text \ Apostol often uses statements like `$P(s_k)$ for all $k$ in an infinite sequence $s_k$ such that $\mathfrak{R}(s_k)\longrightarrow\infty$ as $k\to\infty$'. Instead, we write @{prop "frequently P (Re going_to at_top)"}. This lemma shows that our statement is equivalent to his. \ lemma frequently_going_to_at_top_iff: "frequently P (f going_to (at_top :: real filter)) \ (\g. \n. P (g n) \ filterlim (\n. f (g n)) at_top sequentially)" by (auto intro: frequently_going_toI elim!: frequently_going_to_at_topE) (* END TODO *) lemma surj_bullet_1: "surj (\s::'a::{real_normed_algebra_1, real_inner}. s \ 1)" proof (rule surjI) fix x :: real show "(x *\<^sub>R 1) \ (1 :: 'a) = x" by (simp add: dot_square_norm) qed lemma bullet_1_going_to_at_top_neq_bot [simp]: "((\s::'a::{real_normed_algebra_1, real_inner}. s \ 1) going_to at_top) \ bot" unfolding going_to_def by (rule filtercomap_neq_bot_surj[OF _ surj_bullet_1]) auto lemma fds_abs_converges_altdef: "fds_abs_converges f s \ (\n. fds_nth f n / nat_power n s) abs_summable_on {1..}" by (auto simp add: fds_abs_converges_def abs_summable_on_nat_iff intro!: summable_cong eventually_mono[OF eventually_gt_at_top[of 0]]) lemma fds_abs_converges_altdef': "fds_abs_converges f s \ (\n. fds_nth f n / nat_power n s) abs_summable_on UNIV" by (subst fds_abs_converges_altdef, rule abs_summable_on_cong_neutral) (auto simp: Suc_le_eq) lemma eval_fds_altdef: assumes "fds_abs_converges f s" shows "eval_fds f s = (\\<^sub>an. fds_nth f n / nat_power n s)" proof - have "fds_abs_converges f s \ (\n. fds_nth f n / nat_power n s) abs_summable_on UNIV" unfolding fds_abs_converges_altdef by (intro abs_summable_on_cong_neutral) (auto simp: Suc_le_eq) with assms show ?thesis unfolding eval_fds_def fds_abs_converges_altdef by (intro infsetsum_nat' [symmetric]) simp_all qed lemma multiplicative_function_divide_nat_power: fixes f :: "nat \ 'a :: {nat_power, field}" assumes "multiplicative_function f" shows "multiplicative_function (\n. f n / nat_power n s)" proof interpret f: multiplicative_function f by fact show "f 0 / nat_power 0 s = 0" "f 1 / nat_power 1 s = 1" by simp_all fix a b :: nat assume "a > 1" "b > 1" "coprime a b" thus "f (a * b) / nat_power (a * b) s = f a / nat_power a s * (f b / nat_power b s)" by (simp_all add: f.mult_coprime nat_power_mult_distrib) qed lemma completely_multiplicative_function_divide_nat_power: fixes f :: "nat \ 'a :: {nat_power, field}" assumes "completely_multiplicative_function f" shows "completely_multiplicative_function (\n. f n / nat_power n s)" proof interpret f: completely_multiplicative_function f by fact show "f 0 / nat_power 0 s = 0" "f (Suc 0) / nat_power (Suc 0) s = 1" by simp_all fix a b :: nat assume "a > 1" "b > 1" thus "f (a * b) / nat_power (a * b) s = f a / nat_power a s * (f b / nat_power b s)" by (simp_all add: f.mult nat_power_mult_distrib) qed subsection \Convergence and absolute convergence\ class nat_power_normed_field = nat_power_field + real_normed_field + real_inner + real_algebra_1 + fixes real_power :: "real \ 'a \ 'a" assumes real_power_nat_power: "n > 0 \ real_power (real n) c = nat_power n c" assumes real_power_1_right_aux: "d > 0 \ real_power d 1 = d *\<^sub>R 1" assumes real_power_add: "d > 0 \ real_power d (a + b) = real_power d a * real_power d b" assumes real_power_nonzero [simp]: "d > 0 \ real_power d a \ 0" assumes norm_real_power: "x > 0 \ norm (real_power x c) = x powr (c \ 1)" assumes nat_power_of_real_aux: "nat_power n (x *\<^sub>R 1) = ((real n powr x) *\<^sub>R 1)" assumes has_field_derivative_nat_power_aux: "\x::'a. n > 0 \ LIM y inf_class.inf (Inf (principal ` {S. open S \ x \ S})) (principal (UNIV - {x})). (nat_power n y - nat_power n x - ln (real n) *\<^sub>R nat_power n x * (y - x)) /\<^sub>R norm (y - x) :> Inf (principal ` {S. open S \ 0 \ S})" assumes has_vector_derivative_real_power_aux: "x > 0 \ filterlim (\y. (real_power y c - real_power x (c :: 'a) - (y - x) *\<^sub>R (c * real_power x (c - 1))) /\<^sub>R norm (y - x)) (INF S\{S. open S \ 0 \ S}. principal S) (at x)" assumes norm_nat_power: "n > 0 \ norm (nat_power n y) = real n powr (y \ 1)" begin lemma real_power_diff: "d > 0 \ real_power d (a - b) = real_power d a / real_power d b" using real_power_add[of d b "a - b"] by (simp add: field_simps) end lemma real_power_1_right [simp]: "d > 0 \ real_power d 1 = of_real d" using real_power_1_right_aux[of d] by (simp add: scaleR_conv_of_real) lemma has_vector_derivative_real_power [derivative_intros]: "x > 0 \ ((\y. real_power y c) has_vector_derivative c * real_power x (c - 1)) (at x within A)" by (rule has_vector_derivative_at_within) (insert has_vector_derivative_real_power_aux[of x c], simp add: has_vector_derivative_def has_derivative_def nhds_def bounded_linear_scaleR_left) lemma has_field_derivative_nat_power [derivative_intros]: "n > 0 \ ((\y. nat_power n y) has_field_derivative ln (real n) *\<^sub>R nat_power n x) (at (x :: 'a :: nat_power_normed_field) within A)" by (rule has_field_derivative_at_within) (insert has_field_derivative_nat_power_aux[of n x], simp only: has_field_derivative_def has_derivative_def netlimit_at, simp add: nhds_def at_within_def bounded_linear_mult_right) lemma continuous_on_real_power [continuous_intros]: "A \ {0<..} \ continuous_on A (\x. real_power x s)" by (rule continuous_on_vector_derivative has_vector_derivative_real_power)+ auto instantiation real :: nat_power_normed_field begin definition real_power_real :: "real \ real \ real" where [simp]: "real_power_real = (powr)" instance proof (standard, goal_cases) case (7 n x) hence "((\x. nat_power n x) has_field_derivative ln (real n) *\<^sub>R nat_power n x) (at x)" by (auto intro!: derivative_eq_intros simp: powr_def) thus ?case unfolding has_field_derivative_def netlimit_at has_derivative_def by (simp add: nhds_def at_within_def) next case (8 x c) hence "((\y. real_power y c) has_vector_derivative c * real_power x (c - 1)) (at x)" by (auto intro!: derivative_eq_intros simp: has_field_derivative_iff_has_vector_derivative [symmetric]) thus ?case by (simp add: has_vector_derivative_def has_derivative_def nhds_def) qed (simp_all add: powr_add) end instantiation complex :: nat_power_normed_field begin definition nat_power_complex :: "nat \ complex \ complex" where [simp]: "nat_power_complex n z = of_nat n powr z" definition real_power_complex :: "real \ complex \ complex" where [simp]: "real_power_complex = (\x y. of_real x powr y)" instance proof fix m n :: nat and z :: complex assume "m > 0" "n > 0" thus "nat_power (m * n) z = nat_power m z * nat_power n z" unfolding nat_power_complex_def of_nat_mult by (subst powr_times_real) simp_all next fix n :: nat and z :: complex assume "n > 0" show "norm (nat_power n z) = real n powr (z \ 1)" unfolding nat_power_complex_def using norm_powr_real_powr[of "of_nat n" z] by simp next fix n :: nat and x :: complex assume n: "n > 0" hence "((\x. nat_power n x) has_field_derivative ln (real n) *\<^sub>R nat_power n x) (at x)" by (auto intro!: derivative_eq_intros simp: powr_def scaleR_conv_of_real mult_ac) thus "LIM y inf_class.inf (Inf (principal ` {S. open S \ x \ S})) (principal (UNIV - {x})). (nat_power n y - nat_power n x - ln (real n) *\<^sub>R nat_power n x * (y - x)) /\<^sub>R cmod (y - x) :> (Inf (principal ` {S. open S \ 0 \ S}))" unfolding has_field_derivative_def netlimit_at has_derivative_def by (simp add: nhds_def at_within_def) next fix x :: real and c :: complex assume "x > 0" hence "((\y. real_power y c) has_vector_derivative c * real_power x (c - 1)) (at x)" by (auto intro!: derivative_eq_intros has_vector_derivative_real_field) thus "LIM y at x. (real_power y c - real_power x c - (y - x) *\<^sub>R (c * real_power x (c - 1))) /\<^sub>R norm (y - x) :> INF S\{S. open S \ 0 \ S}. principal S" by (simp add: has_vector_derivative_def has_derivative_def nhds_def) next fix n :: nat and x :: real show "nat_power n (x *\<^sub>R 1 :: complex) = (real n powr x) *\<^sub>R 1" by (simp add: powr_Reals_eq scaleR_conv_of_real) qed (auto simp: powr_def exp_add exp_of_nat_mult [symmetric] algebra_simps scaleR_conv_of_real simp del: Ln_of_nat) end lemma nat_power_of_real [simp]: "nat_power n (of_real x :: 'a :: nat_power_normed_field) = of_real (real n powr x)" using nat_power_of_real_aux[of n x] by (simp add: scaleR_conv_of_real) lemma fds_abs_converges_of_real [simp]: "fds_abs_converges (fds_of_real f) (of_real s :: 'a :: {nat_power_normed_field,banach}) \ fds_abs_converges f s" unfolding fds_abs_converges_def by (subst (1 2) summable_Suc_iff [symmetric]) (simp add: norm_divide norm_nat_power) lemma eval_fds_of_real [simp]: assumes "fds_converges f s" shows "eval_fds (fds_of_real f) (of_real s :: 'a :: {nat_power_normed_field,banach}) = of_real (eval_fds f s)" using assms unfolding eval_fds_def by (auto simp: fds_converges_def suminf_of_real) lemma fds_abs_summable_zeta_iff [simp]: fixes s :: "'a :: {banach, nat_power_normed_field}" shows "fds_abs_converges fds_zeta s \ s \ 1 > (1 :: real)" proof - have "fds_abs_converges fds_zeta s \ summable (\n. real n powr -(s \ 1))" unfolding fds_abs_converges_def by (intro summable_cong always_eventually) (auto simp: norm_divide fds_nth_zeta powr_minus norm_nat_power divide_simps) also have "\ \ s \ 1 > 1" by (simp add: summable_real_powr_iff) finally show ?thesis . qed lemma fds_abs_summable_zeta: "(s :: 'a :: {banach, nat_power_normed_field}) \ 1 > 1 \ fds_abs_converges fds_zeta s" by simp lemma fds_abs_converges_moebius_mu: fixes s :: "'a :: {banach,nat_power_normed_field}" assumes "s \ 1 > 1" shows "fds_abs_converges (fds moebius_mu) s" unfolding fds_abs_converges_def proof (rule summable_comparison_test, intro exI allI impI) fix n :: nat show "norm (norm (fds_nth (fds moebius_mu) n / nat_power n s)) \ real n powr (-s \ 1)" by (auto simp: powr_minus divide_simps abs_moebius_mu_le norm_nat_power norm_divide moebius_mu_def norm_power) next from assms show "summable (\n. real n powr (-s \ 1))" by (simp add: summable_real_powr_iff) qed definition conv_abscissa :: "'a :: {nat_power,banach,real_normed_field, real_inner} fds \ ereal" where "conv_abscissa f = (INF s\{s. fds_converges f s}. ereal (s \ 1))" definition abs_conv_abscissa :: "'a :: {nat_power,banach,real_normed_field, real_inner} fds \ ereal" where "abs_conv_abscissa f = (INF s\{s. fds_abs_converges f s}. ereal (s \ 1))" lemma conv_abscissa_mono: assumes "\s. fds_converges g s \ fds_converges f s" shows "conv_abscissa f \ conv_abscissa g" unfolding conv_abscissa_def by (rule INF_mono) (use assms in auto) lemma abs_conv_abscissa_mono: assumes "\s. fds_abs_converges g s \ fds_abs_converges f s" shows "abs_conv_abscissa f \ abs_conv_abscissa g" unfolding abs_conv_abscissa_def by (rule INF_mono) (use assms in auto) class dirichlet_series = euclidean_space + real_normed_field + nat_power_normed_field + assumes one_in_Basis: "1 \ Basis" instance real :: dirichlet_series by standard simp_all instance complex :: dirichlet_series by standard (simp_all add: Basis_complex_def) context assumes "SORT_CONSTRAINT('a :: dirichlet_series)" begin lemma fds_abs_converges_Re_le: fixes f :: "'a fds" assumes "fds_abs_converges f z" "z \ 1 \ z' \ 1" shows "fds_abs_converges f z'" unfolding fds_abs_converges_def proof (rule summable_comparison_test, intro exI allI impI) fix n :: nat assume n: "n \ 1" thus "norm (norm (fds_nth f n / nat_power n z')) \ norm (fds_nth f n / nat_power n z)" using assms(2) by (simp add: norm_divide norm_nat_power divide_simps powr_mono mult_left_mono) qed (insert assms(1), simp add: fds_abs_converges_def) lemma fds_abs_converges: assumes "s \ 1 > abs_conv_abscissa (f :: 'a fds)" shows "fds_abs_converges f s" proof - from assms obtain s0 where "fds_abs_converges f s0" "s0 \ 1 < s \ 1" by (auto simp: INF_less_iff abs_conv_abscissa_def) with fds_abs_converges_Re_le[OF this(1), of s] this(2) show ?thesis by simp qed lemma fds_abs_diverges: assumes "s \ 1 < abs_conv_abscissa (f :: 'a fds)" shows "\fds_abs_converges f s" proof assume "fds_abs_converges f s" hence "abs_conv_abscissa f \ s \ 1" unfolding abs_conv_abscissa_def by (intro INF_lower) auto with assms show False by simp qed lemma uniformly_Cauchy_eval_fds_aux: fixes s0 :: "'a :: dirichlet_series" assumes bounded: "Bseq (\n. \k\n. fds_nth f k / nat_power k s0)" assumes B: "compact B" "\z. z \ B \ z \ 1 > s0 \ 1" shows "uniformly_Cauchy_on B (\N z. \n\N. fds_nth f n / nat_power n z)" proof (cases "B = {}") case False show ?thesis proof (rule uniformly_Cauchy_onI', goal_cases) case (1 \) define \ where "\ = Inf ((\s. s \ 1) ` B)" have \_le: "s \ 1 \ \" if "s \ B" for s unfolding \_def using that by (intro cInf_lower bounded_inner_imp_bdd_below compact_imp_bounded B) auto have "\ \ ((\s. s \ 1) ` B)" unfolding \_def using B \B \ {}\ by (intro closed_contains_Inf bounded_inner_imp_bdd_below compact_imp_bounded B compact_imp_closed compact_continuous_image continuous_intros) auto with B(2) have \_gt: "\ > s0 \ 1" by auto define \ where "\ = \ - s0 \ 1" have "bounded B" by (rule compact_imp_bounded) fact then obtain norm_B_aux where norm_B_aux: "\s. s \ B \ norm s \ norm_B_aux" by (auto simp: bounded_iff) define norm_B where "norm_B = norm_B_aux + norm s0" from norm_B_aux have norm_B: "norm (s - s0) \ norm_B" if "s \ B" for s using norm_triangle_ineq4[of s s0] norm_B_aux[OF that] by (simp add: norm_B_def) define A where "A = sum_upto (\k. fds_nth f k / nat_power k s0)" from bounded obtain C_aux where C_aux: "\n. norm (\k\n. fds_nth f k / nat_power k s0) \ C_aux" by (auto simp: Bseq_def) define C where "C = max C_aux 1" have C_pos: "C > 0" by (simp add: C_def) have C: "norm (A x) \ C" for x proof - have "A x = (\k\nat \x\. fds_nth f k / nat_power k s0)" unfolding A_def sum_upto_altdef by (intro sum.mono_neutral_left) auto also have "norm \ \ C_aux" by (rule C_aux) also have "\ \ C" by (simp add: C_def) finally show ?thesis . qed have "(\m. 2 * C * (1 + norm_B / \) * real m powr (-\)) \ 0" unfolding \_def using \_gt by (intro tendsto_mult_right_zero tendsto_neg_powr filterlim_real_sequentially) simp_all from order_tendstoD(2)[OF this \\ > 0\] obtain M where M: "\m. m \ M \ 2 * C * (1 + norm_B / \) * real m powr - \ < \" by (auto simp: eventually_at_top_linorder) show ?case proof (intro exI[of _ "max M 1"] ballI allI impI, goal_cases) case (1 s m n) from 1 have s: "s \ 1 > s0 \ 1" using B(2)[of s] by simp have mn: "m \ M" "m < n" "m > 0" "n > 0" using 1 by (simp_all add: ) have "dist (\n\m. fds_nth f n / nat_power n s) (\n\n. fds_nth f n / nat_power n s) = dist (\n\n. fds_nth f n / nat_power n s) (\n\m. fds_nth f n / nat_power n s)" by (simp add: dist_commute) also from 1 have "\ = norm (\k\{..n}-{..m}. fds_nth f k / nat_power k s)" by (subst Groups_Big.sum_diff) (simp_all add: dist_norm) also from 1 have "{..n} - {..m} = real -` {real m<..real n}" by auto also have "(\k\\. fds_nth f k / nat_power k s) = (\k\\. fds_nth f k / nat_power k s0 * real_power (real k) (s0 - s))" (is "_ = ?S") by (intro sum.cong refl) (simp_all add: nat_power_diff real_power_nat_power) also have *: "((\t. A t * ((s0 - s) * real_power t (s0 - s - 1))) has_integral (A (real n) * real_power n (s0 - s) - A (real m) * real_power m (s0 - s) - ?S)) {real m..real n}" (is "(?h has_integral _) _") unfolding A_def using mn by (intro partial_summation_strong[of "{}"]) (auto intro!: derivative_eq_intros continuous_intros) hence "?S = A (real n) * nat_power n (s0 - s) - A (real m) * nat_power m (s0 - s) - integral {real m..real n} ?h" using mn by (simp add: has_integral_iff real_power_nat_power) also have "norm \ \ norm (A (real n) * nat_power n (s0 - s)) + norm (A (real m) * nat_power m (s0 - s)) + norm (integral {real m..real n} ?h)" by (intro order.trans[OF norm_triangle_ineq4] add_right_mono order.refl) also have "norm (A (real n) * nat_power n (s0 - s)) \ C * nat_power m ((s0 - s) \ 1)" using mn \s \ B\ C_pos s by (auto simp: norm_mult norm_nat_power algebra_simps intro!: mult_mono C powr_mono2') also have "norm (A (real m) * nat_power m (s0 - s)) \ C * nat_power m ((s0 - s) \ 1)" using mn by (auto simp: norm_mult norm_nat_power intro!: mult_mono C) also have "norm (integral {real m..real n} ?h) \ integral {real m..real n} (\t. C * (norm (s0 - s) * t powr ((s0 - s) \ 1 - 1)))" proof (intro integral_norm_bound_integral ballI, goal_cases) case 1 with * show ?case by (simp add: has_integral_iff) next case 2 from mn show ?case by (auto intro!: integrable_continuous_real continuous_intros) next case (3 t) thus ?case unfolding norm_mult using C_pos mn by (intro mult_mono C) (auto simp: norm_real_power dot_square_norm algebra_simps) qed also have "\ = C * norm (s0 - s) * integral {real m..real n} (\t. t powr ((s0 - s) \ 1 - 1))" by (simp add: algebra_simps dot_square_norm) also { have "((\t. t powr ((s0 - s) \ 1 - 1)) has_integral (real n powr ((s0 - s) \ 1) / ((s0 - s) \ 1) - real m powr ((s0 - s) \ 1) / ((s0 - s) \ 1))) {m..n}" (is "(?l has_integral ?I) _") using mn s by (intro fundamental_theorem_of_calculus) (auto intro!: derivative_eq_intros simp: has_field_derivative_iff_has_vector_derivative [symmetric] inner_diff_left) hence "integral {real m..real n} ?l = ?I" by (simp add: has_integral_iff) also have "\ \ -(real m powr ((s0 - s) \ 1) / ((s0 - s) \ 1))" using s mn by (simp add: divide_simps inner_diff_left) also have "\ = 1 * (real m powr ((s0 - s) \ 1) / ((s - s0) \ 1))" using s by (simp add: field_simps inner_diff_left) also have "\ \ 2 * (real m powr ((s0 - s) \ 1) / ((s - s0) \ 1))" using mn s by (intro mult_right_mono divide_nonneg_pos) (simp_all add: inner_diff_left) finally have "integral {m..n} ?l \ \" . } hence "C * norm (s0 - s) * integral {real m..real n} (\t. t powr ((s0 - s) \ 1 - 1)) \ C * norm (s0 - s) * (2 * (real m powr ((s0 - s) \ 1) / ((s - s0) \ 1)))" using C_pos mn by (intro mult_mono mult_nonneg_nonneg integral_nonneg integrable_continuous_real continuous_intros) auto also have "C * nat_power m ((s0 - s) \ 1) + C * nat_power m ((s0 - s) \ 1) + \ = 2 * C * nat_power m ((s0 - s) \ 1) * (1 + norm (s - s0) / ((s - s0) \ 1))" by (simp add: algebra_simps norm_minus_commute) also have "\ \ 2 * C * nat_power m (-\) * (1 + norm_B / \)" using C_pos s mn \_le[of s] \s \ B\ \_gt unfolding nat_power_real_def by (intro mult_mono mult_nonneg_nonneg powr_mono frac_le add_mono norm_B) (simp_all add: inner_diff_left \_def) also have "\ = 2 * C * (1 + norm_B / \) * real m powr (-\)" by simp also from \m \ M\ have "\ < \" by (rule M) finally show ?case by - simp_all qed qed qed (auto simp: uniformly_Cauchy_on_def) lemma uniformly_convergent_eval_fds_aux: assumes "Bseq (\n. \k\n. fds_nth f k / nat_power k (s0 :: 'a))" assumes B: "compact B" "\z. z \ B \ z \ 1 > s0 \ 1" shows "uniformly_convergent_on B (\N z. \n\N. fds_nth f n / nat_power n z)" by (rule Cauchy_uniformly_convergent uniformly_Cauchy_eval_fds_aux assms)+ lemma uniformly_convergent_eval_fds_aux': assumes conv: "fds_converges f (s0 :: 'a)" assumes B: "compact B" "\z. z \ B \ z \ 1 > s0 \ 1" shows "uniformly_convergent_on B (\N z. \n\N. fds_nth f n / nat_power n z)" proof (rule uniformly_convergent_eval_fds_aux) from conv have "convergent (\n. \k\n. fds_nth f k / nat_power k s0)" by (simp add: fds_converges_def summable_iff_convergent') thus "Bseq (\n. \k\n. fds_nth f k / nat_power k s0)" by (rule convergent_imp_Bseq) qed (insert assms, auto) lemma bounded_partial_sums_imp_fps_converges: fixes s0 :: "'a :: dirichlet_series" assumes "Bseq (\n. \k\n. fds_nth f k / nat_power k s0)" and "s \ 1 > s0 \ 1" shows "fds_converges f s" proof - have "uniformly_convergent_on {s} (\N z. \n\N. fds_nth f n / nat_power n z)" using assms(2) by (intro uniformly_convergent_eval_fds_aux[OF assms(1)]) auto thus ?thesis by (auto simp: fds_converges_def summable_iff_convergent' dest: uniformly_convergent_imp_convergent) qed theorem fds_converges_Re_le: assumes "fds_converges f (s0 :: 'a)" "s \ 1 > s0 \ 1" shows "fds_converges f s" proof - have "uniformly_convergent_on {s} (\N z. \n\N. fds_nth f n / nat_power n z)" by (rule uniformly_convergent_eval_fds_aux' assms)+ (insert assms(2), auto) then obtain l where "uniform_limit {s} (\N z. \n\N. fds_nth f n / nat_power n z) l at_top" by (auto simp: uniformly_convergent_on_def) from tendsto_uniform_limitI[OF this, of s] have "(\n. fds_nth f n / nat_power n s) sums l s" unfolding sums_def' by (simp add: atLeast0AtMost) thus ?thesis by (simp add: fds_converges_def sums_iff) qed lemma fds_converges: assumes "s \ 1 > conv_abscissa (f :: 'a fds)" shows "fds_converges f s" proof - from assms obtain s0 where "fds_converges f s0" "s0 \ 1 < s \ 1" by (auto simp: INF_less_iff conv_abscissa_def) with fds_converges_Re_le[OF this(1), of s] this(2) show ?thesis by simp qed lemma fds_diverges: assumes "s \ 1 < conv_abscissa (f :: 'a fds)" shows "\fds_converges f s" proof assume "fds_converges f s" hence "conv_abscissa f \ s \ 1" unfolding conv_abscissa_def by (intro INF_lower) auto with assms show False by simp qed theorem fds_converges_imp_abs_converges: assumes "fds_converges (f :: 'a fds) s" "s' \ 1 > s \ 1 + 1" shows "fds_abs_converges f s'" unfolding fds_abs_converges_def proof (rule summable_comparison_test_ev) from assms(2) show "summable (\n. real n powr ((s - s') \ 1))" by (subst summable_real_powr_iff) (simp_all add: inner_diff_left) next from assms(1) have "(\n. fds_nth f n / nat_power n s) \ 0" unfolding fds_converges_def by (rule summable_LIMSEQ_zero) from tendsto_norm[OF this] have "(\n. norm (fds_nth f n / nat_power n s)) \ 0" by simp hence "eventually (\n. norm (fds_nth f n / nat_power n s) < 1) at_top" by (rule order_tendstoD) simp_all thus "eventually (\n. norm (norm (fds_nth f n / nat_power n s')) \ real n powr ((s - s') \ 1)) at_top" proof eventually_elim case (elim n) thus ?case proof (cases "n = 0") case False have "norm (fds_nth f n / nat_power n s') = norm (fds_nth f n) / real n powr (s' \ 1)" using False by (simp add: norm_divide norm_nat_power) also have "\ = norm (fds_nth f n / nat_power n s) / real n powr ((s' - s) \ 1)" using False by (simp add: norm_divide norm_nat_power inner_diff_left powr_diff) also have "\ \ 1 / real n powr ((s' - s) \ 1)" using elim by (intro divide_right_mono elim) simp_all also have "\ = real n powr ((s - s') \ 1)" using False by (simp add: field_simps inner_diff_left powr_diff) finally show ?thesis by simp qed simp_all qed qed lemma conv_le_abs_conv_abscissa: "conv_abscissa f \ abs_conv_abscissa f" unfolding conv_abscissa_def abs_conv_abscissa_def by (intro INF_superset_mono) auto lemma conv_abscissa_PInf_iff: "conv_abscissa f = \ \ (\s. \fds_converges f s)" unfolding conv_abscissa_def by (subst Inf_eq_PInfty) auto lemma conv_abscissa_PInfI [intro]: "(\s. \fds_converges f s) \ conv_abscissa f = \" by (subst conv_abscissa_PInf_iff) auto lemma conv_abscissa_MInf_iff: "conv_abscissa (f :: 'a fds) = -\ \ (\s. fds_converges f s)" proof safe assume *: "\s. fds_converges f s" have "conv_abscissa f \ B" for B :: real using spec[OF *, of "of_real B"] fds_diverges[of "of_real B" f] by (cases "conv_abscissa f \ B") simp_all thus "conv_abscissa f = -\" by (rule ereal_bot) qed (auto intro: fds_converges) lemma conv_abscissa_MInfI [intro]: "(\s. fds_converges (f::'a fds) s) \ conv_abscissa f = -\" by (subst conv_abscissa_MInf_iff) auto lemma abs_conv_abscissa_PInf_iff: "abs_conv_abscissa f = \ \ (\s. \fds_abs_converges f s)" unfolding abs_conv_abscissa_def by (subst Inf_eq_PInfty) auto lemma abs_conv_abscissa_PInfI [intro]: "(\s. \fds_converges f s) \ abs_conv_abscissa f = \" by (subst abs_conv_abscissa_PInf_iff) auto lemma abs_conv_abscissa_MInf_iff: "abs_conv_abscissa (f :: 'a fds) = -\ \ (\s. fds_abs_converges f s)" proof safe assume *: "\s. fds_abs_converges f s" have "abs_conv_abscissa f \ B" for B :: real using spec[OF *, of "of_real B"] fds_abs_diverges[of "of_real B" f] by (cases "abs_conv_abscissa f \ B") simp_all thus "abs_conv_abscissa f = -\" by (rule ereal_bot) qed (auto intro: fds_abs_converges) lemma abs_conv_abscissa_MInfI [intro]: "(\s. fds_abs_converges (f::'a fds) s) \ abs_conv_abscissa f = -\" by (subst abs_conv_abscissa_MInf_iff) auto lemma conv_abscissa_geI: assumes "\c'. ereal c' < c \ \s. s \ 1 = c' \ \fds_converges f s" shows "conv_abscissa (f :: 'a fds) \ c" proof (rule ccontr) assume "\conv_abscissa f \ c" hence "c > conv_abscissa f" by simp from ereal_dense2[OF this] obtain c' where "c > ereal c'" "c' > conv_abscissa f" by auto moreover from assms[OF this(1)] obtain s where "s \ 1 = c'" "\fds_converges f s" by blast ultimately show False using fds_converges[of f s] by auto qed lemma conv_abscissa_leI: assumes "\c'. ereal c' > c \ \s. s \ 1 = c' \ fds_converges f s" shows "conv_abscissa (f :: 'a fds) \ c" proof (rule ccontr) assume "\conv_abscissa f \ c" hence "c < conv_abscissa f" by simp from ereal_dense2[OF this] obtain c' where "c < ereal c'" "c' < conv_abscissa f" by auto moreover from assms[OF this(1)] obtain s where "s \ 1 = c'" "fds_converges f s" by blast ultimately show False using fds_diverges[of s f] by auto qed lemma abs_conv_abscissa_geI: assumes "\c'. ereal c' < c \ \s. s \ 1 = c' \ \fds_abs_converges f s" shows "abs_conv_abscissa (f :: 'a fds) \ c" proof (rule ccontr) assume "\abs_conv_abscissa f \ c" hence "c > abs_conv_abscissa f" by simp from ereal_dense2[OF this] obtain c' where "c > ereal c'" "c' > abs_conv_abscissa f" by auto moreover from assms[OF this(1)] obtain s where "s \ 1 = c'" "\fds_abs_converges f s" by blast ultimately show False using fds_abs_converges[of f s] by auto qed lemma abs_conv_abscissa_leI: assumes "\c'. ereal c' > c \ \s. s \ 1 = c' \ fds_abs_converges f s" shows "abs_conv_abscissa (f :: 'a fds) \ c" proof (rule ccontr) assume "\abs_conv_abscissa f \ c" hence "c < abs_conv_abscissa f" by simp from ereal_dense2[OF this] obtain c' where "c < ereal c'" "c' < abs_conv_abscissa f" by auto moreover from assms[OF this(1)] obtain s where "s \ 1 = c'" "fds_abs_converges f s" by blast ultimately show False using fds_abs_diverges[of s f] by auto qed lemma conv_abscissa_leI_weak: assumes "\x. ereal x > d \ fds_converges f (of_real x)" shows "conv_abscissa (f :: 'a fds) \ d" proof (rule conv_abscissa_leI) fix x assume "d < ereal x" from assms[OF this] show "\s. s \ 1 = x \ fds_converges f s" by (intro exI[of _ "of_real x"]) auto qed lemma abs_conv_abscissa_leI_weak: assumes "\x. ereal x > d \ fds_abs_converges f (of_real x)" shows "abs_conv_abscissa (f :: 'a fds) \ d" proof (rule abs_conv_abscissa_leI) fix x assume "d < ereal x" from assms[OF this] show "\s. s \ 1 = x \ fds_abs_converges f s" by (intro exI[of _ "of_real x"]) auto qed lemma conv_abscissa_truncate [simp]: "conv_abscissa (fds_truncate m (f :: 'a fds)) = -\" by (auto simp: conv_abscissa_MInf_iff) lemma abs_conv_abscissa_truncate [simp]: "abs_conv_abscissa (fds_truncate m (f :: 'a fds)) = -\" by (auto simp: abs_conv_abscissa_MInf_iff) theorem abs_conv_le_conv_abscissa_plus_1: "abs_conv_abscissa (f :: 'a fds) \ conv_abscissa f + 1" proof (rule abs_conv_abscissa_leI) fix c assume less: "conv_abscissa f + 1 < ereal c" define c' where "c' = (if conv_abscissa f = -\ then c - 2 else (c - 1 + real_of_ereal (conv_abscissa f)) / 2)" from less have c': "conv_abscissa f < ereal c' \ c' < c - 1" by (cases "conv_abscissa f") (simp_all add: c'_def field_simps) from c' have "fds_converges f (of_real c')" by (intro fds_converges) (simp_all add: inner_diff_left dot_square_norm) hence "fds_abs_converges f (of_real c)" by (rule fds_converges_imp_abs_converges) (insert c', simp_all) thus "\s. s \ 1 = c \ fds_abs_converges f s" by (intro exI[of _ "of_real c"]) auto qed lemma uniformly_convergent_eval_fds: assumes B: "compact B" "\z. z \ B \ z \ 1 > conv_abscissa (f :: 'a fds)" shows "uniformly_convergent_on B (\N z. \n\N. fds_nth f n / nat_power n z)" proof (cases "B = {}") case False define \ where "\ = Inf ((\s. s \ 1) ` B)" have \_le: "s \ 1 \ \" if "s \ B" for s unfolding \_def using that by (intro cInf_lower bounded_inner_imp_bdd_below compact_imp_bounded B) auto have "\ \ ((\s. s \ 1) ` B)" unfolding \_def using B \B \ {}\ by (intro closed_contains_Inf bounded_inner_imp_bdd_below compact_imp_bounded B compact_imp_closed compact_continuous_image continuous_intros) auto with B(2) have \_gt: "\ > conv_abscissa f" by auto define s where "s = (if conv_abscissa f = -\ then \ - 1 else (\ + real_of_ereal (conv_abscissa f)) / 2)" from \_gt have s: "conv_abscissa f < s \ s < \" by (cases "conv_abscissa f") (auto simp: s_def) show ?thesis using s \compact B\ by (intro uniformly_convergent_eval_fds_aux'[of f "of_real s"] fds_converges) (auto dest: \_le) qed auto corollary uniformly_convergent_eval_fds': assumes B: "compact B" "\z. z \ B \ z \ 1 > conv_abscissa (f :: 'a fds)" shows "uniformly_convergent_on B (\N z. \nN z. \n\N. fds_nth f n / nat_power n z) l at_top" by (auto simp: uniformly_convergent_on_def) also have "(\N z. \n\N. fds_nth f n / nat_power n z) = (\N z. \nN z. \nDerivative of a Dirichlet series\ lemma fds_converges_deriv_aux: assumes conv: "fds_converges f (s0 :: 'a)" and gt: "s \ 1 > s0 \ 1" shows "fds_converges (fds_deriv f) s" proof - have "Cauchy (\n. \k\n. (-ln (real k) *\<^sub>R fds_nth f k) / nat_power k s)" proof (rule CauchyI', goal_cases) case (1 \) define \ where "\ = s \ 1 - s0 \ 1" define \' where "\' = \ / 2" from gt have \_pos: "\ > 0" by (simp add: \_def) define A where "A = sum_upto (\k. fds_nth f k / nat_power k s0)" from conv have "convergent (\n. \k\n. fds_nth f k / nat_power k s0)" by (simp add: fds_converges_def summable_iff_convergent') hence "Bseq (\n. \k\n. fds_nth f k / nat_power k s0)" by (rule convergent_imp_Bseq) then obtain C_aux where C_aux: "\n. norm (\k\n. fds_nth f k / nat_power k s0) \ C_aux" by (auto simp: Bseq_def) define C where "C = max C_aux 1" have C_pos: "C > 0" by (simp add: C_def) have C: "norm (A x) \ C" for x proof - have "A x = (\k\nat \x\. fds_nth f k / nat_power k s0)" unfolding A_def sum_upto_altdef by (intro sum.mono_neutral_left) auto also have "norm \ \ C_aux" by (rule C_aux) also have "\ \ C" by (simp add: C_def) finally show ?thesis . qed define C' where "C' = 2 * C + C * (norm (s0 - s) * (1 + 1 / \) + 1) / \" have "(\m. C' * real m powr (-\')) \ 0" unfolding \'_def using gt \_pos by (intro tendsto_mult_right_zero tendsto_neg_powr filterlim_real_sequentially) simp_all from order_tendstoD(2)[OF this \\ > 0\] obtain M1 where M1: "\m. m \ M1 \ C' * real m powr - \' < \" by (auto simp: eventually_at_top_linorder) have "((\x. ln (real x) / real x powr \') \ 0) at_top" using \_pos by (intro lim_ln_over_power) (simp_all add: \'_def) from order_tendstoD(2)[OF this zero_less_one] eventually_gt_at_top[of "1::nat"] have "eventually (\n. ln (real n) \ n powr \') at_top" by eventually_elim simp_all then obtain M2 where M2: "\n. n \ M2 \ ln (real n) \ n powr \'" by (auto simp: eventually_at_top_linorder) let ?f' = "\k. -ln (real k) *\<^sub>R fds_nth f k" show ?case proof (intro exI[of _ "max (max M1 M2) 1"] allI impI, goal_cases) case (1 m n) hence mn: "m \ M1" "m \ M2" "m > 0" "m < n" by simp_all define g :: "real \ 'a" where "g = (\t. real_power t (s0 - s) * of_real (ln t))" define g' :: "real \ 'a" where "g' = (\t. real_power t (s0 - s - 1) * ((s0 - s) * of_real (ln t) + 1))" define norm_g' :: "real \ real" where "norm_g' = (\t. t powr (-\ - 1) * (norm (s0 - s) * ln t + 1))" define norm_g :: "real \ real" where "norm_g = (\t. -(t powr -\) * (norm (s0 - s) * (\ * ln t + 1) + \) / \^2)" have g_g': "(g has_vector_derivative g' t) (at t)" if "t \ {real m..real n}" for t using mn that by (auto simp: g_def g'_def real_power_diff field_simps real_power_add intro!: derivative_eq_intros) have [continuous_intros]: "continuous_on {real m..real n} g" using mn by (auto simp: g_def intro!: continuous_intros) let ?S = "\k\real -` {real m<..real n}. fds_nth f k / nat_power k s0 * g k" have "dist (\k\m. ?f' k / nat_power k s) (\k\n. ?f' k / nat_power k s) = norm (\k\{..n} - {..m}. fds_nth f k / nat_power k s * of_real (ln (real k)))" using mn by (subst sum_diff) (simp_all add: dist_norm norm_minus_commute sum_negf scaleR_conv_of_real mult_ac) also have "{..n} - {..m} = real -` {real m<..real n}" by auto also have "(\k\\. fds_nth f k / nat_power k s * of_real (ln (real k))) = (\k\\. fds_nth f k / nat_power k s0 * g k)" using mn unfolding g_def by (intro sum.cong refl) (auto simp: real_power_nat_power field_simps nat_power_diff) also have *: "((\t. A t * g' t) has_integral (A (real n) * g n - A (real m) * g m - ?S)) {real m..real n}" (is "(?h has_integral _) _") unfolding A_def using mn by (intro partial_summation_strong[of "{}"]) (auto intro!: g_g' simp: field_simps continuous_intros) hence "?S = A (real n) * g n - A (real m) * g m - integral {real m..real n} ?h" using mn by (simp add: has_integral_iff field_simps) also have "norm \ \ norm (A (real n) * g n) + norm (A (real m) * g m) + norm (integral {real m..real n} ?h)" by (intro order.trans[OF norm_triangle_ineq4] add_right_mono order.refl) also have "norm (A (real n) * g n) \ C * norm (g n)" unfolding norm_mult using mn C_pos by (intro mult_mono C) auto also have "norm (g n) \ n powr -\ * n powr \'" using mn M2[of n] by (simp add: g_def norm_real_power norm_mult \_def inner_diff_left) also have "\ = n powr -\'" using mn by (simp add: \'_def powr_minus field_simps powr_add [symmetric]) also have "norm (A (real m) * g m) \ C * norm (g m)" unfolding norm_mult using mn C_pos by (intro mult_mono C) auto also have "norm (g m) \ m powr -\ * m powr \'" using mn M2[of m] by (simp add: g_def norm_real_power norm_mult \_def inner_diff_left) also have "\ = m powr -\'" using mn by (simp add: \'_def powr_minus field_simps powr_add [symmetric]) also have "C * real n powr - \' \ C * real m powr - \'" using \_pos mn C_pos by (intro mult_left_mono powr_mono2') (simp_all add: \'_def) also have "\ + \ = 2 * \" by simp also have "norm (integral {m..n} ?h) \ integral {m..n} (\t. C * norm_g' t)" proof (intro integral_norm_bound_integral ballI, goal_cases) case 1 with * show ?case by (simp add: has_integral_iff) next case 2 from mn show ?case by (auto intro!: integrable_continuous_real continuous_intros simp: norm_g'_def) next case (3 t) have "norm (g' t) \ norm_g' t" unfolding g'_def norm_g'_def using 3 mn unfolding norm_mult by (intro mult_mono order.trans[OF norm_triangle_ineq]) (auto simp: norm_real_power inner_diff_left dot_square_norm norm_mult \_def intro!: mult_left_mono) thus ?case unfolding norm_mult using C_pos mn by (intro mult_mono C) simp_all qed also have "\ = C * integral {m..n} norm_g'" unfolding norm_g'_def by (simp add: norm_g'_def \_def inner_diff_left) also { have "(norm_g' has_integral (norm_g n - norm_g m)) {m..n}" unfolding norm_g'_def norm_g_def power2_eq_square using mn \_pos by (intro fundamental_theorem_of_calculus) (auto simp: has_field_derivative_iff_has_vector_derivative [symmetric] field_simps powr_diff intro!: derivative_eq_intros) hence "integral {m..n} norm_g' = norm_g n - norm_g m" by (simp add: has_integral_iff) also have "norm_g n \ 0" unfolding norm_g_def using \_pos mn by (intro divide_nonpos_pos mult_nonpos_nonneg add_nonneg_nonneg mult_nonneg_nonneg) simp_all hence "norm_g n - norm_g m \ -norm_g m" by simp also have "\ = real m powr -\ * ln (real m) * (norm (s0 - s)) / \ + real m powr -\ * ((norm (s0 - s) / \ + 1) / \)" using \_pos by (simp add: field_simps norm_g_def power2_eq_square) also { have "ln (real m) \ real m powr \'" using M2[of m] mn by simp also have "real m powr -\ * \ = real m powr -\'" by (simp add: powr_add [symmetric] \'_def) finally have "real m powr -\ * ln (real m) * (norm (s0 - s)) / \ \ \ * (norm (s0 - s)) / \" using \_pos by (intro divide_right_mono mult_right_mono) (simp_all add: mult_left_mono) } also have "real m powr -\ * ((norm (s0 - s) / \ + 1) / \) \ real m powr -\' * ((norm (s0 - s) / \ + 1) / \)" using mn \_pos by (intro mult_right_mono powr_mono) (simp_all add: \'_def) also have "real m powr - \' * norm (s0 - s) / \ + \ = real m powr -\' * (norm (s0 - s) * (1 + 1 / \) + 1) / \" using \_pos by (simp add: field_simps power2_eq_square) finally have "integral {real m..real n} norm_g' \ real m powr - \' * (norm (s0 - s) * (1 + 1 / \) + 1) / \" by - simp_all } also have "2 * (C * m powr - \') + C * (m powr - \' * (norm (s0 - s) * (1 + 1 / \) + 1) / \) = C' * m powr -\'" by (simp add: algebra_simps C'_def) also have "\ < \" using M1[of m] mn by simp finally show ?case using C_pos by - simp_all qed qed from Cauchy_convergent[OF this] show ?thesis by (simp add: summable_iff_convergent' fds_converges_def fds_nth_deriv) qed theorem assumes "s \ 1 > conv_abscissa (f :: 'a fds)" shows fds_converges_deriv: "fds_converges (fds_deriv f) s" and has_field_derivative_eval_fds [derivative_intros]: "(eval_fds f has_field_derivative eval_fds (fds_deriv f) s) (at s within A)" proof - define s1 :: real where "s1 = (if conv_abscissa f = -\ then s \ 1 - 2 else (s \ 1 * 1 / 3 + real_of_ereal (conv_abscissa f) * 2 / 3))" define s2 :: real where "s2 = (if conv_abscissa f = -\ then s \ 1 - 1 else (s \ 1 * 2 / 3 + real_of_ereal (conv_abscissa f) * 1 / 3))" from assms have s: "conv_abscissa f < s1 \ s1 < s2 \ s2 < s \ 1" by (cases "conv_abscissa f") (auto simp: s1_def s2_def field_simps) from s have *: "fds_converges f (of_real s1)" by (intro fds_converges) simp_all thus conv': "fds_converges (fds_deriv f) s" by (rule fds_converges_deriv_aux) (insert s, simp_all) from * have conv: "fds_converges (fds_deriv f) (of_real s2)" by (rule fds_converges_deriv_aux) (insert s, simp_all) define \ :: real where "\ = (s \ 1 - s2) / 2" from s have \_pos: "\ > 0" by (simp add: \_def) have "uniformly_convergent_on (cball s \) (\n s. \k\n. fds_nth (fds_deriv f) k / nat_power k s)" proof (intro uniformly_convergent_eval_fds_aux'[OF conv]) fix s'' :: 'a assume s'': "s'' \ cball s \" have "dist (s \ 1) (s'' \ 1) \ dist s s''" by (intro Euclidean_dist_upper) (simp_all add: one_in_Basis) also from s'' have "\ \ \" by simp finally show "s'' \ 1 > (of_real s2 :: 'a) \ 1" using s by (auto simp: \_def dist_real_def abs_if split: if_splits) qed (insert \_pos, auto) then obtain l where "uniform_limit (cball s \) (\n s. \k\n. fds_nth (fds_deriv f) k / nat_power k s) l at_top" by (auto simp: uniformly_convergent_on_def) also have "(\n s. \k\n. fds_nth (fds_deriv f) k / nat_power k s) = (\n s. \k) (\n s. \k) (\n s. \k cball s \" "s \ interior (cball s \)" using s by (simp_all add: \_def) show "summable (\n. fds_nth f n / nat_power n s)" using assms fds_converges[of f s] by (simp add: fds_converges_def) next fix s' :: 'a and n :: nat show "((\s. fds_nth f n / nat_power n s) has_field_derivative fds_nth (fds_deriv f) n / nat_power n s') (at s' within cball s \)" by (cases "n = 0") (simp, auto intro!: derivative_eq_intros simp: fds_nth_deriv field_simps) qed (auto simp: fds_nth_deriv intro!: derivative_eq_intros) thus "(eval_fds f has_field_derivative eval_fds (fds_deriv f) s) (at s within A)" by (rule has_field_derivative_at_within) qed lemmas has_field_derivative_eval_fds' [derivative_intros] = DERIV_chain2[OF has_field_derivative_eval_fds] lemma continuous_eval_fds [continuous_intros]: assumes "s \ 1 > conv_abscissa f" shows "continuous (at s within A) (eval_fds (f :: 'a :: dirichlet_series fds))" proof - have "isCont (eval_fds f) s" by (rule has_field_derivative_eval_fds DERIV_isCont assms)+ thus ?thesis by (rule continuous_within_subset) auto qed lemma continuous_eval_fds' [continuous_intros]: fixes f :: "'a :: dirichlet_series fds" assumes "continuous (at s within A) g" "g s \ 1 > conv_abscissa f" shows "continuous (at s within A) (\x. eval_fds f (g x))" by (rule continuous_within_compose3[OF _ assms(1)] continuous_intros assms)+ lemma continuous_on_eval_fds [continuous_intros]: fixes f :: "'a :: dirichlet_series fds" assumes "A \ {s. s \ 1 > conv_abscissa f}" shows "continuous_on A (eval_fds f)" by (rule DERIV_continuous_on derivative_intros)+ (insert assms, auto) lemma continuous_on_eval_fds' [continuous_intros]: fixes f :: "'a :: dirichlet_series fds" assumes "continuous_on A g" "g ` A \ {s. s \ 1 > conv_abscissa f}" shows "continuous_on A (\x. eval_fds f (g x))" by (rule continuous_on_compose2[OF continuous_on_eval_fds assms(1)]) (insert assms, auto simp: image_iff) lemma conv_abscissa_deriv_le: fixes f :: "'a fds" shows "conv_abscissa (fds_deriv f) \ conv_abscissa f" proof (rule conv_abscissa_leI) fix c' :: real assume "ereal c' > conv_abscissa f" thus "\s. s \ 1 = c' \ fds_converges (fds_deriv f) s" by (intro exI[of _ "of_real c'"]) (auto simp: fds_converges_deriv) qed lemma abs_conv_abscissa_integral: fixes f :: "'a fds" shows "abs_conv_abscissa (fds_integral a f) = abs_conv_abscissa f" proof (rule antisym) show "abs_conv_abscissa (fds_integral a f) \ abs_conv_abscissa f" proof (rule abs_conv_abscissa_leI, goal_cases) case (1 c) have "fds_abs_converges (fds_integral a f) (of_real c)" unfolding fds_abs_converges_def proof (rule summable_comparison_test_ev) from 1 have "fds_abs_converges f (of_real c)" by (intro fds_abs_converges) auto thus "summable (\n. norm (fds_nth f n / nat_power n (of_real c)))" by (simp add: fds_abs_converges_def) next show "\\<^sub>F n in sequentially. norm (norm (fds_nth (fds_integral a f) n / nat_power n (of_real c))) \ norm (fds_nth f n / nat_power n (of_real c))" using eventually_gt_at_top[of 3] proof eventually_elim case (elim n) from elim and exp_le have "ln (exp 1) \ ln (real n)" by (subst ln_le_cancel_iff) auto hence "1 * norm (fds_nth f n) \ ln (real n) * norm (fds_nth f n)" by (intro mult_right_mono) auto with elim show ?case by (simp add: norm_divide norm_nat_power fds_integral_def field_simps) qed qed thus ?case by (intro exI[of _ "of_real c"]) auto qed next show "abs_conv_abscissa f \ abs_conv_abscissa (fds_integral a f)" (is "_ \ ?s0") proof (cases "abs_conv_abscissa (fds_integral a f) = \") case False show ?thesis proof (rule abs_conv_abscissa_leI) fix c :: real define \ where "\ = (if ?s0 = -\ then 1 else (c - real_of_ereal ?s0) / 2)" assume "ereal c > ?s0" with False have \: "\ > 0" "c - \ > ?s0" by (cases ?s0; force simp: \_def field_simps)+ have "fds_abs_converges f (of_real c)" unfolding fds_abs_converges_def proof (rule summable_comparison_test_ev) from \ have "fds_abs_converges (fds_integral a f) (of_real (c - \))" by (intro fds_abs_converges) (auto simp: algebra_simps) thus "summable (\n. norm (fds_nth (fds_integral a f) n / nat_power n (of_real (c - \))))" by (simp add: fds_abs_converges_def) next have "\\<^sub>F n in at_top. ln (real n) / real n powr \ < 1" by (rule order_tendstoD lim_ln_over_power \\ > 0\ zero_less_one)+ thus "\\<^sub>F n in sequentially. norm (norm (fds_nth f n / nat_power n (of_real c))) \ norm (fds_nth (fds_integral a f) n / nat_power n (of_real (c - \)))" using eventually_gt_at_top[of 1] proof eventually_elim case (elim n) hence "ln (real n) * norm (fds_nth f n) \ real n powr \ * norm (fds_nth f n)" by (intro mult_right_mono) auto with elim show ?case by (simp add: norm_divide norm_nat_power field_simps powr_diff inner_diff_left fds_integral_def) qed qed thus "\s. s \ 1 = c \ fds_abs_converges f s" by (intro exI[of _ "of_real c"]) auto qed qed auto qed lemma abs_conv_abscissa_ln: "abs_conv_abscissa (fds_ln l (f :: 'a :: dirichlet_series fds)) = abs_conv_abscissa (fds_deriv f / f)" by (simp add: fds_ln_def abs_conv_abscissa_integral) lemma abs_conv_abscissa_deriv: fixes f :: "'a fds" shows "abs_conv_abscissa (fds_deriv f) = abs_conv_abscissa f" proof - have "abs_conv_abscissa (fds_deriv f) = abs_conv_abscissa (fds_integral (fds_nth f 1) (fds_deriv f))" by (rule abs_conv_abscissa_integral [symmetric]) also have "fds_integral (fds_nth f 1) (fds_deriv f) = f" by (rule fds_integral_fds_deriv) finally show ?thesis . qed lemma abs_conv_abscissa_higher_deriv: "abs_conv_abscissa ((fds_deriv ^^ n) f) = abs_conv_abscissa (f :: 'a :: dirichlet_series fds)" by (induction n) (simp_all add: abs_conv_abscissa_deriv) lemma conv_abscissa_higher_deriv_le: "conv_abscissa ((fds_deriv ^^ n) f) \ conv_abscissa (f :: 'a :: dirichlet_series fds)" by (induction n) (auto intro: order.trans[OF conv_abscissa_deriv_le]) lemma abs_conv_abscissa_restrict: "abs_conv_abscissa (fds_subseries P f) \ abs_conv_abscissa f" by (rule abs_conv_abscissa_mono) auto lemma eval_fds_deriv: fixes f :: "'a fds" assumes "s \ 1 > conv_abscissa f" shows "eval_fds (fds_deriv f) s = deriv (eval_fds f) s" by (intro DERIV_imp_deriv [symmetric] derivative_intros assms) lemma eval_fds_higher_deriv: assumes "(s :: 'a :: dirichlet_series) \ 1 > conv_abscissa f" shows "eval_fds ((fds_deriv ^^ n) f) s = (deriv ^^ n) (eval_fds f) s" using assms proof (induction n arbitrary: f s) case (Suc n f s) have ev: "eventually (\s. s \ {s. s \ 1 > conv_abscissa f}) (nhds s)" using Suc.prems open_halfspace_gt[of _ "1::'a"] by (intro eventually_nhds_in_open, cases "conv_abscissa f") (auto simp: open_halfspace_gt inner_commute) have "eval_fds ((fds_deriv ^^ Suc n) f) s = eval_fds ((fds_deriv ^^ n) (fds_deriv f)) s" by (subst funpow_Suc_right) simp also have "\ = (deriv ^^ n) (eval_fds (fds_deriv f)) s" by (intro Suc.IH le_less_trans[OF conv_abscissa_deriv_le] Suc.prems) also have "\ = (deriv ^^ n) (deriv (eval_fds f)) s" by (intro higher_deriv_cong_ev refl eventually_mono[OF ev] eval_fds_deriv) auto also have "\ = (deriv ^^ Suc n) (eval_fds f) s" by (subst funpow_Suc_right) simp finally show ?case . qed auto end subsection \Multiplication of two series\ lemma fixes f g :: "nat \ 'a :: {banach, real_normed_field, second_countable_topology, nat_power}" fixes s :: 'a assumes [simp]: "f 0 = 0" "g 0 = 0" assumes summable: "summable (\n. norm (f n / nat_power n s))" "summable (\n. norm (g n / nat_power n s))" shows summable_dirichlet_prod: "summable (\n. norm (dirichlet_prod f g n / nat_power n s))" and suminf_dirichlet_prod: "(\n. dirichlet_prod f g n / nat_power n s) = (\n. f n / nat_power n s) * (\n. g n / nat_power n s)" proof - have summable': "(\n. f n / nat_power n s) abs_summable_on A" "(\n. g n / nat_power n s) abs_summable_on A" for A by ((rule abs_summable_on_subset[OF _ subset_UNIV], insert summable, simp add: abs_summable_on_nat_iff'); fail)+ have f_g: "f a / nat_power a s * (g b / nat_power b s) = f a * g b / nat_power (a * b) s" for a b by (cases "a * b = 0") (auto simp: nat_power_mult_distrib) have eq: "(\\<^sub>a(m, n)\{(m, n). m * n = x}. f m * g n / nat_power x s) = dirichlet_prod f g x / nat_power x s" for x :: nat proof (cases "x > 0") case False hence "(\\<^sub>a(m,n) | m * n = x. f m * g n / nat_power x s) = (\\<^sub>a(m,n) | m * n = x. 0)" by (intro infsetsum_cong) auto with False show ?thesis by simp next case True from finite_divisors_nat'[OF this] show ?thesis by (simp add: dirichlet_prod_altdef2 case_prod_unfold sum_divide_distrib) qed have "(\(m,n). (f m / nat_power m s) * (g n / nat_power n s)) abs_summable_on UNIV \ UNIV" using summable' by (intro abs_summable_on_product) auto also have "?this \ (\(m,n). f m * g n / nat_power (m*n) s) abs_summable_on UNIV" using f_g by (intro abs_summable_on_cong) auto also have "\ \ (\(x,(m,n)). f m * g n / nat_power (m*n) s) abs_summable_on (SIGMA x:UNIV. {(m,n). m * n = x})" unfolding case_prod_unfold by (rule abs_summable_on_reindex_bij_betw [symmetric]) (auto simp: bij_betw_def inj_on_def image_iff) also have "\ \ (\(x,(m,n)). f m * g n / nat_power x s) abs_summable_on (SIGMA x:UNIV. {(m,n). m * n = x})" by (intro abs_summable_on_cong) auto finally have summable'': \ . from abs_summable_on_Sigma_project1'[OF this] show summable''': "summable (\n. norm (dirichlet_prod f g n / nat_power n s))" by (simp add: eq abs_summable_on_nat_iff') have "(\n. f n / nat_power n s) * (\n. g n / nat_power n s) = (\\<^sub>an. f n / nat_power n s) * (\\<^sub>an. g n / nat_power n s)" using summable' by (simp add: infsetsum_nat') also have "\ = (\\<^sub>a(m,n). (f m / nat_power m s) * (g n / nat_power n s))" using summable' by (subst infsetsum_product [symmetric]) simp_all also have "\ = (\\<^sub>a(m,n). f m * g n / nat_power (m * n) s)" using f_g by (intro infsetsum_cong refl) auto also have "\ = (\\<^sub>a(x,(m,n))\(SIGMA x:UNIV. {(m,n). m * n = x}). f m * g n / nat_power (m * n) s)" unfolding case_prod_unfold by (rule infsetsum_reindex_bij_betw [symmetric]) (auto simp: bij_betw_def inj_on_def image_iff) also have "\ = (\\<^sub>a(x,(m,n))\(SIGMA x:UNIV. {(m,n). m * n = x}). f m * g n / nat_power x s)" by (intro infsetsum_cong refl) (auto simp: case_prod_unfold) also have "\ = (\\<^sub>ax. dirichlet_prod f g x / nat_power x s)" (is "_ = infsetsum ?T _") using summable'' by (subst infsetsum_Sigma) (auto simp: eq) also have "\ = (\x. dirichlet_prod f g x / nat_power x s)" using summable''' by (intro infsetsum_nat') (simp_all add: abs_summable_on_nat_iff') finally show "\ = (\n. f n / nat_power n s) * (\n. g n / nat_power n s)" .. qed lemma fixes f g :: "nat \ real" fixes s :: real assumes "f 0 = 0" "g 0 = 0" assumes summable: "summable (\n. norm (f n / real n powr s))" "summable (\n. norm (g n / real n powr s))" shows summable_dirichlet_prod_real: "summable (\n. norm (dirichlet_prod f g n / real n powr s))" and suminf_dirichlet_prod_real: "(\n. dirichlet_prod f g n / real n powr s) = (\n. f n / nat_power n s) * (\n. g n / real n powr s)" using summable_dirichlet_prod[of f g s] suminf_dirichlet_prod[of f g s] assms by simp_all lemma fds_abs_converges_mult: fixes s :: "'a :: {nat_power, real_normed_field, banach, second_countable_topology}" assumes "fds_abs_converges f s" "fds_abs_converges g s" shows "fds_abs_converges (f * g) s" using summable_dirichlet_prod[OF _ _ assms[unfolded fds_abs_converges_def]] by (simp add: fds_abs_converges_def fds_nth_mult) lemma fds_abs_converges_power: fixes s :: "'a :: {nat_power, real_normed_field, banach, second_countable_topology}" shows "fds_abs_converges f s \ fds_abs_converges (f ^ n) s" by (induction n) (auto intro!: fds_abs_converges_mult) lemma fds_abs_converges_prod: fixes s :: "'a :: {nat_power, real_normed_field, banach, second_countable_topology}" shows "(\x. x \ A \ fds_abs_converges (f x) s) \ fds_abs_converges (prod f A) s" by (induction A rule: infinite_finite_induct) (auto intro!: fds_abs_converges_mult) lemma abs_conv_abscissa_mult_le: "abs_conv_abscissa (f * g :: 'a :: dirichlet_series fds) \ max (abs_conv_abscissa f) (abs_conv_abscissa g)" proof (rule abs_conv_abscissa_leI, goal_cases) case (1 c') thus ?case by (auto intro!: exI[of _ "of_real c'"] fds_abs_converges_mult intro: fds_abs_converges) qed lemma abs_conv_abscissa_mult_leI: "abs_conv_abscissa (f :: 'a :: dirichlet_series fds) \ d \ abs_conv_abscissa g \ d \ abs_conv_abscissa (f * g) \ d" using abs_conv_abscissa_mult_le[of f g] by (auto simp add: le_max_iff_disj) lemma abs_conv_abscissa_shift [simp]: "abs_conv_abscissa (fds_shift c f) = abs_conv_abscissa (f :: 'a :: dirichlet_series fds) + c \ 1" proof - have "abs_conv_abscissa (fds_shift c f) \ abs_conv_abscissa f + c \ 1" for c :: 'a and f proof (rule abs_conv_abscissa_leI) fix d assume "abs_conv_abscissa f + c \ 1 < ereal d" hence "abs_conv_abscissa f < ereal (d - c \ 1)" by (cases "abs_conv_abscissa f") auto hence "fds_abs_converges (fds_shift c f) (of_real d)" by (auto intro!: fds_abs_converges_shift fds_abs_converges simp: algebra_simps) thus "\s. s \ 1 = d \ fds_abs_converges (fds_shift c f) s" by (auto intro!: exI[of _ "of_real d"]) qed note * = this[of c f] this[of "-c" "fds_shift c f"] show ?thesis by (cases "abs_conv_abscissa (fds_shift c f)"; cases "abs_conv_abscissa f") (insert *, auto intro!: antisym) qed lemma eval_fds_mult: fixes s :: "'a :: {nat_power, real_normed_field, banach, second_countable_topology}" assumes "fds_abs_converges f s" "fds_abs_converges g s" shows "eval_fds (f * g) s = eval_fds f s * eval_fds g s" using suminf_dirichlet_prod[OF _ _ assms[unfolded fds_abs_converges_def]] by (simp_all add: eval_fds_def fds_nth_mult) lemma eval_fds_power: fixes s :: "'a :: {nat_power, real_normed_field, banach, second_countable_topology}" assumes "fds_abs_converges f s" shows "eval_fds (f ^ n) s = eval_fds f s ^ n" using assms by (induction n) (simp_all add: eval_fds_mult fds_abs_converges_power) lemma eval_fds_prod: fixes s :: "'a :: {nat_power, real_normed_field, banach, second_countable_topology}" assumes "(\x. x \ A \ fds_abs_converges (f x) s)" shows "eval_fds (prod f A) s = (\x\A. eval_fds (f x) s)" using assms by (induction A rule: infinite_finite_induct) (auto simp: eval_fds_mult fds_abs_converges_prod) lemma eval_fds_inverse: fixes s :: "'a :: {nat_power, real_normed_field, banach, second_countable_topology}" assumes "fds_abs_converges f s" "fds_abs_converges (inverse f) s" "fds_nth f 1 \ 0" shows "eval_fds (inverse f) s = inverse (eval_fds f s)" proof - have "eval_fds (inverse f * f) s = eval_fds (inverse f) s * eval_fds f s" by (intro eval_fds_mult assms) also have "inverse f * f = 1" by (intro fds_left_inverse assms) also have "eval_fds 1 s = 1" by simp finally show ?thesis by (auto simp: divide_simps) qed lemma eval_fds_integral_has_field_derivative: fixes s :: "'a :: dirichlet_series" assumes "ereal (s \ 1) > abs_conv_abscissa f" assumes "fds_nth f 1 = 0" shows "(eval_fds (fds_integral c f) has_field_derivative eval_fds f s) (at s)" proof - have "conv_abscissa (fds_integral c f) \ abs_conv_abscissa (fds_integral c f)" by (rule conv_le_abs_conv_abscissa) also from assms have "\ < ereal (s \ 1)" by (simp add: abs_conv_abscissa_integral) finally have "(eval_fds (fds_integral c f) has_field_derivative eval_fds (fds_deriv (fds_integral c f)) s) (at s)" by (intro derivative_eq_intros) auto also from assms have "fds_deriv (fds_integral c f) = f" by simp finally show ?thesis . qed lemma holomorphic_fds_eval [holomorphic_intros]: "A \ {z. Re z > conv_abscissa f} \ eval_fds f holomorphic_on A" unfolding holomorphic_on_def field_differentiable_def by (rule ballI exI derivative_intros)+ auto lemma analytic_fds_eval [holomorphic_intros]: assumes "A \ {z. Re z > conv_abscissa f}" shows "eval_fds f analytic_on A" proof - have "eval_fds f analytic_on {z. Re z > conv_abscissa f}" proof (subst analytic_on_open) show "open {z. Re z > conv_abscissa f}" by (cases "conv_abscissa f") (simp_all add: open_halfspace_Re_gt) qed (intro holomorphic_intros, simp_all) from analytic_on_subset[OF this assms] show ?thesis . qed lemma conv_abscissa_0 [simp]: "conv_abscissa (0 :: 'a :: dirichlet_series fds) = -\" by (auto simp: conv_abscissa_MInf_iff) lemma abs_conv_abscissa_0 [simp]: "abs_conv_abscissa (0 :: 'a :: dirichlet_series fds) = -\" by (auto simp: abs_conv_abscissa_MInf_iff) lemma conv_abscissa_1 [simp]: "conv_abscissa (1 :: 'a :: dirichlet_series fds) = -\" by (auto simp: conv_abscissa_MInf_iff) lemma abs_conv_abscissa_1 [simp]: "abs_conv_abscissa (1 :: 'a :: dirichlet_series fds) = -\" by (auto simp: abs_conv_abscissa_MInf_iff) lemma conv_abscissa_const [simp]: "conv_abscissa (fds_const (c :: 'a :: dirichlet_series)) = -\" by (auto simp: conv_abscissa_MInf_iff) lemma abs_conv_abscissa_const [simp]: "abs_conv_abscissa (fds_const (c :: 'a :: dirichlet_series)) = -\" by (auto simp: abs_conv_abscissa_MInf_iff) lemma conv_abscissa_numeral [simp]: "conv_abscissa (numeral n :: 'a :: dirichlet_series fds) = -\" by (auto simp: numeral_fds) lemma abs_conv_abscissa_numeral [simp]: "abs_conv_abscissa (numeral n :: 'a :: dirichlet_series fds) = -\" by (auto simp: numeral_fds) lemma abs_conv_abscissa_power_le: "abs_conv_abscissa (f ^ n :: 'a :: dirichlet_series fds) \ abs_conv_abscissa f" by (induction n) (auto intro!: order.trans[OF abs_conv_abscissa_mult_le]) lemma abs_conv_abscissa_power_leI: "abs_conv_abscissa (f :: 'a :: dirichlet_series fds) \ d \ abs_conv_abscissa (f ^ n) \ d" by (rule order.trans[OF abs_conv_abscissa_power_le]) lemma abs_conv_abscissa_prod_le: assumes "\x. x \ A \ abs_conv_abscissa (f x :: 'a :: dirichlet_series fds) \ d" shows "abs_conv_abscissa (prod f A) \ d" using assms by (induction A rule: infinite_finite_induct) (auto intro!: abs_conv_abscissa_mult_leI) lemma conv_abscissa_add_le: "conv_abscissa (f + g :: 'a :: dirichlet_series fds) \ max (conv_abscissa f) (conv_abscissa g)" by (rule conv_abscissa_leI_weak) (auto intro!: fds_converges_add intro: fds_converges) lemma conv_abscissa_add_leI: "conv_abscissa (f :: 'a :: dirichlet_series fds) \ d \ conv_abscissa g \ d \ conv_abscissa (f + g) \ d" using conv_abscissa_add_le[of f g] by (auto simp: le_max_iff_disj) lemma conv_abscissa_sum_leI: assumes "\x. x \ A \ conv_abscissa (f x :: 'a :: dirichlet_series fds) \ d" shows "conv_abscissa (sum f A) \ d" using assms by (induction A rule: infinite_finite_induct) (auto intro!: conv_abscissa_add_leI) lemma abs_conv_abscissa_add_le: "abs_conv_abscissa (f + g :: 'a :: dirichlet_series fds) \ max (abs_conv_abscissa f) (abs_conv_abscissa g)" by (rule abs_conv_abscissa_leI_weak) (auto intro!: fds_abs_converges_add intro: fds_abs_converges) lemma abs_conv_abscissa_add_leI: "abs_conv_abscissa (f :: 'a :: dirichlet_series fds) \ d \ abs_conv_abscissa g \ d \ abs_conv_abscissa (f + g) \ d" using abs_conv_abscissa_add_le[of f g] by (auto simp: le_max_iff_disj) lemma abs_conv_abscissa_sum_leI: assumes "\x. x \ A \ abs_conv_abscissa (f x :: 'a :: dirichlet_series fds) \ d" shows "abs_conv_abscissa (sum f A) \ d" using assms by (induction A rule: infinite_finite_induct) (auto intro!: abs_conv_abscissa_add_leI) lemma fds_converges_cmult_left [intro]: assumes "fds_converges f s" shows "fds_converges (fds_const c * f) s" proof - from assms have "summable (\n. c * (fds_nth f n / nat_power n s))" by (intro summable_mult) (auto simp: fds_converges_def) thus ?thesis by (simp add: fds_converges_def mult_ac) qed lemma fds_converges_cmult_right [intro]: assumes "fds_converges f s" shows "fds_converges (f * fds_const c) s" using fds_converges_cmult_left[OF assms] by (simp add: mult_ac) lemma conv_abscissa_cmult_left [simp]: fixes c :: "'a :: dirichlet_series" assumes "c \ 0" shows "conv_abscissa (fds_const c * f) = conv_abscissa f" proof - have "fds_converges (fds_const c * f) s \ fds_converges f s" for s proof assume "fds_converges (fds_const c * f) s" hence "fds_converges (fds_const (inverse c) * (fds_const c * f)) s" by (rule fds_converges_cmult_left) also have "fds_const (inverse c) * (fds_const c * f) = fds_const (inverse c * c) * f" by simp also have "inverse c * c = 1" using assms by simp finally show "fds_converges f s" by simp qed auto thus ?thesis by (simp add: conv_abscissa_def) qed lemma conv_abscissa_cmult_right [simp]: fixes c :: "'a :: dirichlet_series" assumes "c \ 0" shows "conv_abscissa (f * fds_const c) = conv_abscissa f" using assms by (subst mult.commute) auto lemma abs_conv_abscissa_cmult: fixes c :: "'a :: dirichlet_series" assumes "c \ 0" shows "abs_conv_abscissa (fds_const c * f) = abs_conv_abscissa f" proof (intro antisym) have "abs_conv_abscissa (fds_const (inverse c) * (fds_const c * f)) \ abs_conv_abscissa (fds_const c * f)" using abs_conv_abscissa_mult_le[of "fds_const (inverse c)" "fds_const c * f"] by (auto simp: max_def) also have "fds_const (inverse c) * (fds_const c * f) = fds_const (inverse c * c) * f" by (simp add: mult_ac) also have "inverse c * c = 1" using assms by simp finally show "abs_conv_abscissa f \ abs_conv_abscissa (fds_const c * f)" by simp qed (insert abs_conv_abscissa_mult_le[of "fds_const c" f], auto simp: max_def) lemma conv_abscissa_minus [simp]: fixes f :: "'a :: dirichlet_series fds" shows "conv_abscissa (-f) = conv_abscissa f" using conv_abscissa_cmult_left[of "-1" f] by simp lemma abs_conv_abscissa_minus [simp]: fixes f :: "'a :: dirichlet_series fds" shows "abs_conv_abscissa (-f) = abs_conv_abscissa f" using abs_conv_abscissa_cmult[of "-1" f] by simp lemma conv_abscissa_diff_le: "conv_abscissa (f - g :: 'a :: dirichlet_series fds) \ max (conv_abscissa f) (conv_abscissa g)" using conv_abscissa_add_le[of f "-g"] by simp lemma conv_abscissa_diff_leI: "conv_abscissa (f :: 'a :: dirichlet_series fds) \ d \ conv_abscissa g \ d \ conv_abscissa (f - g) \ d" using conv_abscissa_add_le[of f "-g"] by (auto simp: le_max_iff_disj) lemma abs_conv_abscissa_diff_le: "abs_conv_abscissa (f - g :: 'a :: dirichlet_series fds) \ max (abs_conv_abscissa f) (abs_conv_abscissa g)" using abs_conv_abscissa_add_le[of f "-g"] by simp lemma abs_conv_abscissa_diff_leI: "abs_conv_abscissa (f :: 'a :: dirichlet_series fds) \ d \ abs_conv_abscissa g \ d \ abs_conv_abscissa (f - g) \ d" using abs_conv_abscissa_add_le[of f "-g"] by (auto simp: le_max_iff_disj) lemmas eval_fds_integral_has_field_derivative' [derivative_intros] = DERIV_chain'[OF _ eval_fds_integral_has_field_derivative] lemma abs_conv_abscissa_completely_multiplicative_log_deriv: fixes f :: "'a :: dirichlet_series fds" assumes "completely_multiplicative_function (fds_nth f)" "fds_nth f 1 \ 0" shows "abs_conv_abscissa (fds_deriv f / f) \ abs_conv_abscissa f" proof - have "fds_deriv f = - fds (\n. fds_nth f n * mangoldt n) * f" using assms by (subst completely_multiplicative_fds_deriv') simp_all also have "\ / f = - fds (\n. fds_nth f n * mangoldt n) * (f * inverse f)" by (simp add: divide_fds_def) also have "f * inverse f = 1" using assms by (intro fds_right_inverse) finally have "fds_deriv f / f = - fds (\n. fds_nth f n * mangoldt n)" by simp also have "abs_conv_abscissa \ = abs_conv_abscissa (fds (\n. fds_nth f n * mangoldt n))" (is "_ = abs_conv_abscissa ?f") by (rule abs_conv_abscissa_minus) also have "\ \ abs_conv_abscissa f" proof (rule abs_conv_abscissa_leI, goal_cases) case (1 c) have "fds_abs_converges ?f (of_real c)" unfolding fds_abs_converges_def proof (rule summable_comparison_test_ev) from 1 have "fds_abs_converges (fds_deriv f) (of_real c)" by (intro fds_abs_converges) (auto simp: abs_conv_abscissa_deriv) thus "summable (\n. \ln (real n)\ * norm (fds_nth f n) / norm (nat_power n (of_real c :: 'a)))" by (simp add: fds_abs_converges_def fds_deriv_def fds_nth_fds' scaleR_conv_of_real powr_minus norm_mult norm_divide norm_nat_power) next show "\\<^sub>F n in sequentially. norm (norm (fds_nth (fds (\n. fds_nth f n * mangoldt n)) n / nat_power n (of_real c))) \ \ln (real n)\ * norm (fds_nth f n) / norm (nat_power n (of_real c) :: 'a)" using eventually_gt_at_top[of 0] proof eventually_elim case (elim n) have "norm (norm (fds_nth (fds (\n. fds_nth f n * mangoldt n)) n / nat_power n (of_real c))) = norm (fds_nth f n) * mangoldt n / real n powr c" using elim by (simp add: fds_nth_fds' norm_mult norm_divide norm_nat_power abs_mult mangoldt_nonneg) also have "\ \ norm (fds_nth f n) * ln n / real n powr c" using elim by (intro mult_left_mono divide_right_mono mangoldt_le) (simp_all add: mangoldt_def) finally show ?case using elim by (simp add: norm_nat_power algebra_simps) qed qed thus ?case by (intro exI[of _ "of_real c"]) auto qed finally show ?thesis . qed subsection \Uniqueness\ context assumes "SORT_CONSTRAINT ('a :: dirichlet_series)" begin lemma norm_dirichlet_series_cutoff_le: assumes "fds_abs_converges f (s0 :: 'a)" "N > 0" "s \ 1 \ c" "c \ s0 \ 1" shows "summable (\n. fds_nth f (n + N) / nat_power (n + N) s)" "summable (\n. norm (fds_nth f (n + N)) / nat_power (n + N) c)" and "norm (\n. fds_nth f (n + N) / nat_power (n + N) s) \ (\n. norm (fds_nth f (n + N)) / nat_power (n + N) c) / nat_power N (s \ 1 - c)" proof - from assms have "fds_abs_converges f (of_real c)" using fds_abs_converges_Re_le[of f s0 "of_real c"] by auto hence "summable (\n. norm (fds_nth f (n + N) / nat_power (n + N) (of_real c)))" unfolding fds_abs_converges_def by (rule summable_ignore_initial_segment) also have "?this \ summable (\n. norm (fds_nth f (n + N)) / nat_power (n + N) c)" by (intro summable_cong eventually_mono[OF eventually_gt_at_top[of "0::nat"]]) (auto simp: norm_divide norm_nat_power) finally show *: "summable (\n. norm (fds_nth f (n + N)) / nat_power (n + N) c)" . from assms have "fds_abs_converges f s" using fds_abs_converges_Re_le[of f s0 s] by auto hence **: "summable (\n. norm (fds_nth f (n + N) / nat_power (n + N) s))" unfolding fds_abs_converges_def by (rule summable_ignore_initial_segment) thus "summable (\n. fds_nth f (n + N) / nat_power (n + N) s)" by (rule summable_norm_cancel) have "norm (\n. fds_nth f (n + N) / nat_power (n + N) s) \ (\n. norm (fds_nth f (n + N) / nat_power (n + N) s))" by (intro summable_norm **) also have "\ \ (\n. norm (fds_nth f (n + N)) / nat_power (n + N) c / nat_power N (s \ 1 - c))" proof (intro suminf_le * ** summable_divide allI) fix n :: nat have "real N powr (s \ 1 - c) \ real (n + N) powr (s \ 1 - c)" using assms by (intro powr_mono2) simp_all also have "real (n + N) powr c * \ = real (n + N) powr (s \ 1)" by (simp add: powr_diff) finally have "norm (fds_nth f (n + N)) / real (n + N) powr (s \ 1) \ norm (fds_nth f (n + N)) / (real (n + N) powr c * real N powr (s \ 1 - c))" using \N > 0\ by (intro divide_left_mono) (simp_all add: mult_left_mono) thus "norm (fds_nth f (n + N) / nat_power (n + N) s) \ norm (fds_nth f (n + N)) / nat_power (n + N) c / nat_power N (s \ 1 - c)" using \N > 0\ by (simp add: norm_divide norm_nat_power ) qed also have "\ = (\n. norm (fds_nth f (n + N)) / nat_power (n + N) c) / nat_power N (s \ 1 - c)" using * by (rule suminf_divide) finally show "norm (\n. fds_nth f (n + N) / nat_power (n + N) s) \ \" . qed lemma eval_fds_zeroD_aux: fixes h :: "'a fds" assumes conv: "fds_abs_converges h (s0 :: 'a)" assumes freq: "frequently (\s. eval_fds h s = 0) ((\s. s \ 1) going_to at_top)" shows "h = 0" proof (rule ccontr) assume "h \ 0" hence ex: "\n>0. fds_nth h n \ 0" by (auto simp: fds_eq_iff) define N :: nat where "N = (LEAST n. n > 0 \ fds_nth h n \ 0)" have N: "N > 0" "fds_nth h N \ 0" using LeastI_ex[OF ex, folded N_def] by auto have less_N: "fds_nth h n = 0" if "n < N" for n using Least_le[of "\n. n > 0 \ fds_nth h n \ 0" n, folded N_def] that by (cases "n = 0") (auto simp: not_less) define c where "c = s0 \ 1" define remainder where "remainder = (\s. (\n. fds_nth h (n + Suc N) / nat_power (n + Suc N) s))" define A where "A = (\n. norm (fds_nth h (n + Suc N)) / nat_power (n + Suc N) c) * nat_power (Suc N) c" have eq: "fds_nth h N = nat_power N s * eval_fds h s - nat_power N s * remainder s" if "s \ 1 \ c" for s :: 'a proof - from conv and that have conv': "fds_abs_converges h s" unfolding c_def by (rule fds_abs_converges_Re_le) hence conv'': "fds_converges h s" by blast from conv'' have "(\n. fds_nth h n / nat_power n s) sums eval_fds h s" by (simp add: fds_converges_iff) hence "(\n. fds_nth h (n + Suc N) / nat_power (n + Suc N) s) sums (eval_fds h s - (\nnn = fds_nth h N / nat_power N s" by (subst sum.delta) auto finally show ?thesis unfolding remainder_def using \N > 0\ by (auto simp: sums_iff field_simps) qed have remainder_bound: "norm (remainder s) \ A / real (Suc N) powr (s \ 1)" if "s \ 1 \ c" for s :: 'a proof - note * = norm_dirichlet_series_cutoff_le[of h s0 "Suc N" c s, folded remainder_def] have "norm (remainder s) \ (\n. norm (fds_nth h (n + Suc N)) / nat_power (n + Suc N) c) / nat_power (Suc N) (s \ 1 - c)" using that assms unfolding remainder_def by (intro *) (simp_all add: c_def) also have "\ = A / real (Suc N) powr (s \ 1)" by (simp add: A_def powr_diff) finally show ?thesis . qed from freq have "\c. \s. s \ 1 \ c \ eval_fds h s = 0" unfolding frequently_def by (auto simp: eventually_going_to_at_top_linorder) hence "\k. \s. s \ 1 \ real k \ eval_fds h s = 0" by blast then obtain S where S: "\k. S k \ 1 \ real k \ eval_fds h (S k) = 0" by metis have S_limit: "filterlim (\k. S k \ 1) at_top sequentially" by (rule filterlim_at_top_mono[OF filterlim_real_sequentially]) (use S in auto) have "eventually (\k. real k \ c) sequentially" by real_asymp hence "eventually (\k. norm (fds_nth h N) \ (real N / real (Suc N)) powr (S k \ 1) * A) sequentially" proof eventually_elim case (elim k) hence "norm (fds_nth h N) = real N powr (S k \ 1) * norm (remainder (S k))" (is "_ = _ * ?X") using \N > 0\ S[of k] eq[of "S k"] by (auto simp: norm_mult norm_nat_power c_def) also have "norm (remainder (S k)) \ A / real (Suc N) powr (S k \ 1)" using elim S[of k] by (intro remainder_bound) (simp_all add: c_def) finally show ?case using N by (simp add: mult_left_mono powr_divide field_simps del: of_nat_Suc) qed moreover have "((\k. (real N / real (Suc N)) powr (S k \ 1) * A) \ 0) sequentially" by (rule filterlim_compose[OF _ S_limit]) (use \N > 0\ in real_asymp) ultimately have "((\_. fds_nth h N) \ 0) sequentially" by (rule Lim_null_comparison) hence "fds_nth h N = 0" by (simp add: tendsto_const_iff) with \fds_nth h N \ 0\ show False by contradiction qed lemma eval_fds_zeroD: fixes h :: "'a fds" assumes conv: "conv_abscissa h < \" assumes freq: "frequently (\s. eval_fds h s = 0) ((\s. s \ 1) going_to at_top)" shows "h = 0" proof - have [simp]: "2 \ (1 :: 'a) = 2" using of_real_inner_1[of 2] unfolding of_real_numeral by simp from conv obtain s where "fds_converges h s" by auto hence "fds_abs_converges h (s + 2)" by (rule fds_converges_imp_abs_converges) (auto simp: algebra_simps) from this assms(2-) show ?thesis by (rule eval_fds_zeroD_aux) qed lemma eval_fds_eqD: fixes f g :: "'a fds" assumes conv: "conv_abscissa f < \" "conv_abscissa g < \" assumes eq: "frequently (\s. eval_fds f s = eval_fds g s) ((\s. s \ 1) going_to at_top)" shows "f = g" proof - have conv': "conv_abscissa (f - g) < \" using assms by (intro le_less_trans[OF conv_abscissa_diff_le]) (auto simp: max_def) have "max (conv_abscissa f) (conv_abscissa g) < \" using conv by (auto simp: max_def) from ereal_dense2[OF this] obtain c where c: "max (conv_abscissa f) (conv_abscissa g) < ereal c" by auto (* TODO: something like "frequently_elim" would be great here *) have "frequently (\s. eval_fds f s = eval_fds g s \ s \ 1 \ c) ((\s. s \ 1) going_to at_top)" using eq by (rule frequently_eventually_frequently) auto hence *: "frequently (\s. eval_fds (f - g) s = 0) ((\s. s \ 1) going_to at_top)" proof (rule frequently_mono [rotated], safe, goal_cases) case (1 s) thus ?case using c by (subst eval_fds_diff) (auto intro!: fds_converges intro: less_le_trans) qed have "f - g = 0" by (rule eval_fds_zeroD fds_abs_converges_diff assms * conv')+ thus ?thesis by simp qed end subsection \Limit at infinity\ lemma eval_fds_at_top_tail_bound: fixes f :: "'a :: dirichlet_series fds" assumes c: "ereal c > abs_conv_abscissa f" defines "B \ (\n. norm (fds_nth f (n+2)) / real (n+2) powr c) * 2 powr c" assumes s: "s \ 1 \ c" shows "norm (eval_fds f s - fds_nth f 1) \ B / 2 powr (s \ 1)" proof - from c have "fds_abs_converges f (of_real c)" by (intro fds_abs_converges) simp_all also have "?this \ summable (\n. norm (fds_nth f n) / real n powr c)" unfolding fds_abs_converges_def by (intro summable_cong eventually_mono[OF eventually_gt_at_top[of "0::nat"]]) (auto simp: norm_divide norm_nat_power norm_powr_real_powr) finally have summable_c: \ . note c also from s have "ereal c \ ereal (s \ 1)" by simp finally have "fds_abs_converges f s" by (intro fds_abs_converges) auto hence summable: "summable (\n. norm (fds_nth f n / nat_power n s))" by (simp add: fds_abs_converges_def) from summable_norm_cancel[OF this] have "(\n. fds_nth f n / nat_power n s) sums eval_fds f s" by (simp add: eval_fds_def sums_iff) from sums_split_initial_segment[OF this, of "Suc (Suc 0)"] have "norm (eval_fds f s - fds_nth f 1) = norm (\n. fds_nth f (n+2) / nat_power (n+2) s)" by (auto simp: sums_iff) also have "\ \ (\n. norm (fds_nth f (n+2) / nat_power (n+2) s))" by (intro summable_norm summable_ignore_initial_segment summable) also have "\ \ (\n. norm (fds_nth f (n+2)) / real (n+2) powr c / 2 powr (s \ 1 - c))" proof (intro suminf_le allI) fix n :: nat have "norm (fds_nth f (n + 2) / nat_power (n + 2) s) = norm (fds_nth f (n + 2)) / real (n+2) powr c / real (n+2) powr (s \ 1 - c)" by (simp add: field_simps powr_diff norm_divide norm_nat_power) also have "\ \ norm (fds_nth f (n + 2)) / real (n+2) powr c / 2 powr (s \ 1 - c)" using s by (intro divide_left_mono divide_nonneg_pos powr_mono2 mult_pos_pos) simp_all finally show "norm (fds_nth f (n + 2) / nat_power (n + 2) s) \ \" . qed (intro summable_ignore_initial_segment summable summable_divide summable_c)+ also have "\ = (\n. norm (fds_nth f (n+2)) / real (n+2) powr c) / 2 powr (s \ 1 - c)" by (intro suminf_divide summable_ignore_initial_segment summable_c) also have "\ = B / 2 powr (s \ 1)" by (simp add: B_def powr_diff) finally show ?thesis . qed lemma tendsto_eval_fds_Re_at_top: assumes "conv_abscissa (f :: 'a :: dirichlet_series fds) \ \" assumes lim: "filterlim (\x. S x \ 1) at_top F" shows "((\x. eval_fds f (S x)) \ fds_nth f 1) F" proof - from assms(1) have "abs_conv_abscissa f < \" using abs_conv_le_conv_abscissa_plus_1[of f] by auto from ereal_dense2[OF this] obtain c where c: "abs_conv_abscissa f < ereal c" by auto define B where "B = (\n. norm (fds_nth f (n+2)) / real (n+2) powr c) * 2 powr c" have *: "norm (eval_fds f s - fds_nth f 1) \ B / 2 powr (s \ 1)" if s: "s \ 1 \ c" for s using eval_fds_at_top_tail_bound[of f c s] that c by (simp add: B_def) moreover from lim have "eventually (\x. S x \ 1 \ c) F" by (auto simp: filterlim_at_top) ultimately have "eventually (\x. norm (eval_fds f (S x) - fds_nth f 1) \ B / 2 powr (S x \ 1)) F" by (auto elim!: eventually_mono) moreover have "((\x. B / 2 powr (S x \ 1)) \ 0) F" using filterlim_tendsto_pos_mult_at_top[OF tendsto_const[of "ln 2"] _ lim] by (intro real_tendsto_divide_at_top[OF tendsto_const]) (auto simp: powr_def mult_ac intro!: filterlim_compose[OF exp_at_top]) ultimately have "((\x. eval_fds f (S x) - fds_nth f 1) \ 0) F" by (rule Lim_null_comparison) thus ?thesis by (subst (asm) Lim_null [symmetric]) qed lemma tendsto_eval_fds_Re_at_top': assumes "conv_abscissa (f :: complex fds) \ \" shows "uniform_limit UNIV (\\ t. eval_fds f (of_real \ + of_real t * \) ) (\_ .fds_nth f 1) at_top" proof - from assms(1) have "abs_conv_abscissa f < \" using abs_conv_le_conv_abscissa_plus_1[of f] by auto from ereal_dense2[OF this] obtain c where c: "abs_conv_abscissa f < ereal c" by auto define B where "B \ (\n. norm (fds_nth f (n+2)) / real (n+2) powr c) * 2 powr c" show ?thesis unfolding uniform_limit_iff proof safe fix \ :: real assume "\ > 0" hence "eventually (\\. B / 2 powr \ < \) at_top" by real_asymp thus "eventually (\\. \t\UNIV. dist (eval_fds f (of_real \ + of_real t * \)) (fds_nth f 1) < \) at_top" using eventually_ge_at_top[of c] proof eventually_elim case (elim \) show ?case proof fix t :: real have "dist (eval_fds f (of_real \ + of_real t * \)) (fds_nth f 1) \ B / 2 powr \" using eval_fds_at_top_tail_bound[of f c "of_real \ + of_real t * \"] elim c by (simp add: dist_norm B_def) also have "\ < \" by fact finally show "dist (eval_fds f (of_real \ + of_real t * \)) (fds_nth f 1) < \" . qed qed qed qed lemma tendsto_eval_fds_Re_going_to_at_top: assumes "conv_abscissa (f :: 'a :: dirichlet_series fds) \ \" shows "((\s. eval_fds f s) \ fds_nth f 1) ((\s. s \ 1) going_to at_top)" using assms by (rule tendsto_eval_fds_Re_at_top) auto lemma tendsto_eval_fds_Re_going_to_at_top': assumes "conv_abscissa (f :: complex fds) \ \" shows "((\s. eval_fds f s) \ fds_nth f 1) (Re going_to at_top)" using assms by (rule tendsto_eval_fds_Re_at_top) auto text \ Any Dirichlet series that is not identically zero and does not diverge everywhere has a half-plane in which it converges and is non-zero. \ theorem fds_nonzero_halfplane_exists: fixes f :: "'a :: dirichlet_series fds" assumes "conv_abscissa f < \" "f \ 0" shows "eventually (\s. fds_converges f s \ eval_fds f s \ 0) ((\s. s \ 1) going_to at_top)" proof - from ereal_dense2[OF assms(1)] obtain c where c: "conv_abscissa f < ereal c" by auto have "eventually (\s::'a. s \ 1 > c) ((\s. s \ 1) going_to at_top)" using eventually_gt_at_top[of c] by auto hence "eventually (\s. fds_converges f s) ((\s. s \ 1) going_to at_top)" by eventually_elim (use c in \auto intro!: fds_converges simp: less_le_trans\) moreover have "eventually (\s. eval_fds f s \ 0) ((\s. s \ 1) going_to at_top)" using eval_fds_zeroD[OF assms(1)] assms(2) by (auto simp: frequently_def) ultimately show ?thesis by (rule eventually_conj) qed subsection \Normed series\ lemma fds_converges_norm_iff [simp]: fixes s :: "'a :: {nat_power_normed_field,banach}" shows "fds_converges (fds_norm f) (s \ 1) \ fds_abs_converges f s" unfolding fds_converges_def fds_abs_converges_def by (rule summable_cong [OF eventually_mono[OF eventually_gt_at_top[of 0]]]) (simp add: fds_abs_converges_def fds_norm_def fds_nth_fds' norm_divide norm_nat_power) lemma fds_abs_converges_norm_iff [simp]: fixes s :: "'a :: {nat_power_normed_field,banach}" shows "fds_abs_converges (fds_norm f) (s \ 1) \ fds_abs_converges f s" unfolding fds_abs_converges_def by (rule summable_cong [OF eventually_mono[OF eventually_gt_at_top[of 0]]]) (simp add: fds_abs_converges_def fds_norm_def fds_nth_fds' norm_divide norm_nat_power) lemma fds_converges_norm_iff': fixes f :: "'a :: {nat_power_normed_field,banach} fds" shows "fds_converges (fds_norm f) s \ fds_abs_converges f (of_real s)" unfolding fds_converges_def fds_abs_converges_def by (rule summable_cong [OF eventually_mono[OF eventually_gt_at_top[of 0]]]) (simp add: fds_abs_converges_def fds_norm_def fds_nth_fds' norm_divide norm_nat_power) lemma fds_abs_converges_norm_iff': fixes f :: "'a :: {nat_power_normed_field,banach} fds" shows "fds_abs_converges (fds_norm f) s \ fds_abs_converges f (of_real s)" unfolding fds_abs_converges_def by (rule summable_cong [OF eventually_mono[OF eventually_gt_at_top[of 0]]]) (simp add: fds_abs_converges_def fds_norm_def fds_nth_fds' norm_divide norm_nat_power) lemma abs_conv_abscissa_norm [simp]: fixes f :: "'a :: dirichlet_series fds" shows "abs_conv_abscissa (fds_norm f) = abs_conv_abscissa f" proof (rule antisym) show "abs_conv_abscissa f \ abs_conv_abscissa (fds_norm f)" proof (rule abs_conv_abscissa_leI_weak) fix x assume "abs_conv_abscissa (fds_norm f) < ereal x" hence "fds_abs_converges (fds_norm f) (of_real x)" by (intro fds_abs_converges) auto thus "fds_abs_converges f (of_real x)" by (simp add: fds_abs_converges_norm_iff') qed qed (auto intro!: abs_conv_abscissa_leI_weak simp: fds_abs_converges_norm_iff' fds_abs_converges) lemma conv_abscissa_norm [simp]: fixes f :: "'a :: dirichlet_series fds" shows "conv_abscissa (fds_norm f) = abs_conv_abscissa f" proof (rule antisym) show "abs_conv_abscissa f \ conv_abscissa (fds_norm f)" proof (rule abs_conv_abscissa_leI_weak) fix x assume "conv_abscissa (fds_norm f) < ereal x" hence "fds_converges (fds_norm f) (of_real x)" by (intro fds_converges) auto thus "fds_abs_converges f (of_real x)" by (simp add: fds_converges_norm_iff') qed qed (auto intro!: conv_abscissa_leI_weak simp: fds_abs_converges) lemma fixes f g :: "'a :: dirichlet_series fds" assumes "fds_abs_converges (fds_norm f) s" "fds_abs_converges (fds_norm g) s" shows fds_abs_converges_norm_mult: "fds_abs_converges (fds_norm (f * g)) s" and eval_fds_norm_mult_le: "eval_fds (fds_norm (f * g)) s \ eval_fds (fds_norm f) s * eval_fds (fds_norm g) s" proof - show conv: "fds_abs_converges (fds_norm (f * g)) s" unfolding fds_abs_converges_def proof (rule summable_comparison_test_ev) have "fds_abs_converges (fds_norm f * fds_norm g) s" by (rule fds_abs_converges_mult assms)+ thus "summable (\n. norm (fds_nth (fds_norm f * fds_norm g) n) / nat_power n s)" by (simp add: fds_abs_converges_def) qed (auto intro!: always_eventually divide_right_mono order.trans[OF fds_nth_norm_mult_le] simp: norm_divide) have conv': "fds_abs_converges (fds_norm f * fds_norm g) s" by (intro fds_abs_converges_mult assms) hence "eval_fds (fds_norm (f * g)) s \ eval_fds (fds_norm f * fds_norm g) s" using conv unfolding eval_fds_def fds_abs_converges_def norm_divide by (intro suminf_le allI divide_right_mono) (simp_all add: norm_mult fds_nth_norm_mult_le) also have "\ = eval_fds (fds_norm f) s * eval_fds (fds_norm g) s" by (intro eval_fds_mult assms) finally show "eval_fds (fds_norm (f * g)) s \ eval_fds (fds_norm f) s * eval_fds (fds_norm g) s" . qed lemma eval_fds_norm_nonneg: assumes "fds_abs_converges (fds_norm f) s" shows "eval_fds (fds_norm f) s \ 0" using assms unfolding eval_fds_def fds_abs_converges_def by (intro suminf_nonneg) auto lemma fixes f :: "'a :: dirichlet_series fds" assumes "fds_abs_converges (fds_norm f) s" shows fds_abs_converges_norm_power: "fds_abs_converges (fds_norm (f ^ n)) s" and eval_fds_norm_power_le: "eval_fds (fds_norm (f ^ n)) s \ eval_fds (fds_norm f) s ^ n" proof - show *: "fds_abs_converges (fds_norm (f ^ n)) s" for n by (induction n) (auto intro!: fds_abs_converges_norm_mult assms) show "eval_fds (fds_norm (f ^ n)) s \ eval_fds (fds_norm f) s ^ n" by (induction n) (auto intro!: order.trans[OF eval_fds_norm_mult_le] assms * mult_left_mono eval_fds_norm_nonneg) qed subsection \Logarithms of Dirichlet series\ (* TODO: Move ? *) lemma eventually_gt_ereal_at_top: "c \ \ \ eventually (\x. ereal x > c) at_top" by (cases c) auto lemma eval_fds_log_deriv: fixes s :: "'a :: dirichlet_series" assumes "fds_nth f 1 \ 0" "s \ 1 > abs_conv_abscissa f" "s \ 1 > abs_conv_abscissa (fds_deriv f / f)" assumes "eval_fds f s \ 0" shows "eval_fds (fds_deriv f / f) s = eval_fds (fds_deriv f) s / eval_fds f s" proof - have "eval_fds (fds_deriv f / f * f) s = eval_fds (fds_deriv f / f) s * eval_fds f s" using assms by (intro eval_fds_mult fds_abs_converges) auto also have "fds_deriv f / f * f = fds_deriv f * (f * inverse f)" by (simp add: divide_fds_def algebra_simps) also have "f * inverse f = 1" using assms by (intro fds_right_inverse) finally show ?thesis using assms by simp qed text \ Given a sufficiently nice absolutely convergent Dirichlet series that converges to some function $f(s)$ and a holomorphic branch of $\ln f(s)$, we can construct a Dirichlet series that absolutely converges to that logarithm. \ lemma eval_fds_ln: fixes s0 :: ereal assumes nz: "\s. Re s > s0 \ eval_fds f s \ 0" "fds_nth f 1 \ 0" assumes l: "exp l = fds_nth f 1" "((g \ of_real) \ l) at_top" assumes g: "\s. Re s > s0 \ exp (g s) = eval_fds f s" assumes holo_g: "g holomorphic_on {s. Re s > s0}" assumes "ereal (Re s) > s0" assumes "s0 \ abs_conv_abscissa f" and "s0 \ abs_conv_abscissa (fds_deriv f / f)" shows "eval_fds (fds_ln l f) s = g s" proof - let ?s0 = "abs_conv_abscissa f" and ?s1 = "abs_conv_abscissa (inverse f)" let ?h = "\s. eval_fds (fds_ln l f) s - g s" let ?A = "{s. Re s > s0}" have open_A: "open ?A" by (cases s0) (auto simp: open_halfspace_Re_gt) have "conv_abscissa f \ abs_conv_abscissa f" by (rule conv_le_abs_conv_abscissa) moreover from assms have "\ \ \" by auto ultimately have "conv_abscissa f \ \" by auto have "conv_abscissa (fds_ln l f) \ abs_conv_abscissa (fds_ln l f)" by (rule conv_le_abs_conv_abscissa) also have "\ \ abs_conv_abscissa (fds_deriv f / f)" unfolding fds_ln_def by (simp add: abs_conv_abscissa_integral) finally have "conv_abscissa (fds_ln l f) \ \" using assms by (auto simp: max_def abs_conv_abscissa_deriv split: if_splits) have deriv_g [derivative_intros]: "(g has_field_derivative eval_fds (fds_deriv f) s / eval_fds f s) (at s within B)" if s: "Re s > s0" for s B proof - have "conv_abscissa f \ abs_conv_abscissa f" by (rule conv_le_abs_conv_abscissa) also have "\ \ s0" using assms by simp also have "\ < Re s" by fact finally have s': "Re s > conv_abscissa f" . have deriv_g: "(g has_field_derivative deriv g s) (at s)" using holomorphic_derivI[OF holo_g open_A, of s] s by (auto simp: at_within_open[OF _ open_A]) have "((\s. exp (g s)) has_field_derivative eval_fds f s * deriv g s) (at s)" (is ?P) by (rule derivative_eq_intros deriv_g s)+ (insert s, simp_all add: g) also from s have ev: "eventually (\t. t \ ?A) (nhds s)" by (intro eventually_nhds_in_open open_A) auto have "?P \ (eval_fds f has_field_derivative eval_fds f s * deriv g s) (at s)" by (intro DERIV_cong_ev refl eventually_mono[OF ev]) (auto simp: g) finally have "(eval_fds f has_field_derivative eval_fds f s * deriv g s) (at s)" . moreover have "(eval_fds f has_field_derivative eval_fds (fds_deriv f) s) (at s)" using s' assms by (intro derivative_intros) auto ultimately have "eval_fds f s * deriv g s = eval_fds (fds_deriv f) s" by (rule DERIV_unique) hence "deriv g s = eval_fds (fds_deriv f) s / eval_fds f s" using s nz by (simp add: field_simps) with deriv_g show ?thesis by (auto intro: has_field_derivative_at_within) qed have "\c. \z\{z. Re z > s0}. ?h z = c" proof (rule has_field_derivative_zero_constant, goal_cases) case 1 show ?case using convex_halfspace_gt[of _ "1::complex"] by (cases s0) auto next case (2 z) have "conv_abscissa (fds_ln l f) \ abs_conv_abscissa (fds_ln l f)" by (rule conv_le_abs_conv_abscissa) also have "\ \ abs_conv_abscissa (fds_deriv f / f)" by (simp add: abs_conv_abscissa_ln) also have "\ < Re z" using 2 assms by (auto simp: abs_conv_abscissa_deriv) finally have s1: "conv_abscissa (fds_ln l f) < ereal (Re z)" . have "conv_abscissa f \ abs_conv_abscissa f" by (rule conv_le_abs_conv_abscissa) also have "\ < Re z" using 2 assms by auto finally have s2: "conv_abscissa f < ereal (Re z)" . from l have "fds_nth f 1 \ 0" by auto with 2 assms have *: "eval_fds (fds_deriv f / f) z = eval_fds (fds_deriv f) z / (eval_fds f z)" by (auto simp: eval_fds_log_deriv) have "eval_fds f z \ 0" using 2 assms by auto show ?case using s1 s2 2 nz by (auto intro!: derivative_eq_intros simp: * field_simps) qed then obtain c where c: "\z. Re z > s0 \ ?h z = c" by blast have "(at_top :: real filter) \ bot" by simp moreover from assms have "s0 \ \" by auto have "eventually (\x. c = (?h \ of_real) x) at_top" using eventually_gt_ereal_at_top[OF \s0 \ \\] by eventually_elim (simp add: c) hence "((?h \ of_real) \ c) at_top" by (force intro: Lim_transform_eventually) moreover have "((?h \ of_real) \ fds_nth (fds_ln l f) 1 - l) at_top" using \conv_abscissa (fds_ln l f) \ \\ and l unfolding o_def by (intro tendsto_intros tendsto_eval_fds_Re_at_top) (auto simp: filterlim_ident) ultimately have "c = fds_nth (fds_ln l f) 1 - l" by (rule tendsto_unique) with c[OF \Re s > s0\] and l and nz show ?thesis by (simp add: exp_minus field_simps) qed text \ Less explicitly: For a sufficiently nice absolutely convergent Dirichlet series converging to a function $f(s)$, the formal logarithm absolutely converges to some logarithm of $f(s)$. \ lemma eval_fds_ln': fixes s0 :: ereal assumes "ereal (Re s) > s0" assumes "s0 \ abs_conv_abscissa f" and "s0 \ abs_conv_abscissa (fds_deriv f / f)" and nz: "\s. Re s > s0 \ eval_fds f s \ 0" "fds_nth f 1 \ 0" assumes l: "exp l = fds_nth f 1" shows "exp (eval_fds (fds_ln l f) s) = eval_fds f s" proof - let ?s0 = "abs_conv_abscissa f" and ?s1 = "abs_conv_abscissa (inverse f)" let ?h = "\s. eval_fds f s * exp (-eval_fds (fds_ln l f) s)" have "conv_abscissa f \ abs_conv_abscissa f" by (rule conv_le_abs_conv_abscissa) moreover from assms have "\ \ \" by auto ultimately have "conv_abscissa f \ \" by auto have "conv_abscissa (fds_ln l f) \ abs_conv_abscissa (fds_ln l f)" by (rule conv_le_abs_conv_abscissa) also have "\ \ abs_conv_abscissa (fds_deriv f / f)" unfolding fds_ln_def by (simp add: abs_conv_abscissa_integral) finally have "conv_abscissa (fds_ln l f) \ \" using assms by (auto simp: max_def abs_conv_abscissa_deriv split: if_splits) have "\c. \z\{z. Re z > s0}. ?h z = c" proof (rule has_field_derivative_zero_constant, goal_cases) case 1 show ?case using convex_halfspace_gt[of _ "1::complex"] by (cases s0) auto next case (2 z) have "conv_abscissa (fds_ln l f) \ abs_conv_abscissa (fds_ln l f)" by (rule conv_le_abs_conv_abscissa) also have "\ \ abs_conv_abscissa (fds_deriv f / f)" unfolding fds_ln_def by (simp add: abs_conv_abscissa_integral) also have "\ < Re z" using 2 assms by (auto simp: abs_conv_abscissa_deriv) finally have s1: "conv_abscissa (fds_ln l f) < ereal (Re z)" . have "conv_abscissa f \ abs_conv_abscissa f" by (rule conv_le_abs_conv_abscissa) also have "\ < Re z" using 2 assms by auto finally have s2: "conv_abscissa f < ereal (Re z)" . from l have "fds_nth f 1 \ 0" by auto with 2 assms have *: "eval_fds (fds_deriv f / f) z = eval_fds (fds_deriv f) z / (eval_fds f z)" by (subst eval_fds_log_deriv) auto have "eval_fds f z \ 0" using 2 assms by auto thus ?case using s1 s2 by (auto intro!: derivative_eq_intros simp: *) qed then obtain c where c: "\z. Re z > s0 \ ?h z = c" by blast have "(at_top :: real filter) \ bot" by simp moreover from assms have "s0 \ \" by auto have "eventually (\x. c = (?h \ of_real) x) at_top" using eventually_gt_ereal_at_top[OF \s0 \ \\] by eventually_elim (simp add: c) hence "((?h \ of_real) \ c) at_top" by (force intro: Lim_transform_eventually) moreover have "((?h \ of_real) \ fds_nth f 1 * exp (-fds_nth (fds_ln l f) 1)) at_top" unfolding o_def using \conv_abscissa (fds_ln l f) \ \\ and \conv_abscissa f \ \\ by (intro tendsto_intros tendsto_eval_fds_Re_at_top) (auto simp: filterlim_ident) ultimately have "c = fds_nth f 1 * exp (-fds_nth (fds_ln l f) 1)" by (rule tendsto_unique) with c[OF \Re s > s0\] and l and nz show ?thesis by (simp add: exp_minus field_simps) qed lemma fds_ln_completely_multiplicative: fixes f :: "'a :: dirichlet_series fds" assumes "completely_multiplicative_function (fds_nth f)" assumes "fds_nth f 1 \ 0" shows "fds_ln l f = fds (\n. if n = 1 then l else fds_nth f n * mangoldt n /\<^sub>R ln n)" proof - have "fds_ln l f = fds_integral l (fds_deriv f / f)" by (simp add: fds_ln_def) also have "fds_deriv f = -fds (\n. fds_nth f n * mangoldt n) * f" by (intro completely_multiplicative_fds_deriv' assms) also have "\ / f = -fds (\n. fds_nth f n * mangoldt n) * (f * inverse f)" by (simp add: divide_fds_def) also from assms have "f * inverse f = 1" by (simp add: fds_right_inverse) also have "fds_integral l (- fds (\n. fds_nth f n * mangoldt n) * 1) = fds (\n. if n = 1 then l else fds_nth f n * mangoldt n /\<^sub>R ln n)" by (simp add: fds_integral_def cong: if_cong) finally show ?thesis . qed lemma eval_fds_ln_completely_multiplicative_strong: fixes s :: "'a :: dirichlet_series" and l :: 'a and f :: "'a fds" and g :: "nat \ 'a" defines "h \ fds (\n. fds_nth (fds_ln l f) n * g n)" assumes "fds_abs_converges h s" assumes "completely_multiplicative_function (fds_nth f)" and "fds_nth f 1 \ 0" shows "(\(p,k). (fds_nth f p / nat_power p s) ^ Suc k * g (p ^ Suc k) / of_nat (Suc k)) abs_summable_on ({p. prime p} \ UNIV)" (is ?th1) and "eval_fds h s = l * g 1 + (\\<^sub>a(p, k)\{p. prime p}\UNIV. (fds_nth f p / nat_power p s) ^ Suc k * g (p ^ Suc k) / of_nat (Suc k))" (is ?th2) proof - let ?P = "{p::nat. prime p}" interpret f: completely_multiplicative_function "fds_nth f" by fact from assms have *: "(\n. fds_nth h n / nat_power n s) abs_summable_on UNIV" by (auto simp: abs_summable_on_nat_iff' fds_abs_converges_def) have eq: "h = fds (\n. if n = 1 then l * g 1 else fds_nth f n * g n * mangoldt n /\<^sub>R ln (real n))" using fds_ln_completely_multiplicative [OF assms(3), of l] by (simp add: h_def fds_eq_iff) note * also have "(\n. fds_nth h n / nat_power n s) abs_summable_on UNIV \ (\x. if x = Suc 0 then l * g 1 else fds_nth f x * g x * mangoldt x /\<^sub>R ln (real x) / nat_power x s) abs_summable_on {1} \ Collect primepow" using eq by (intro abs_summable_on_cong_neutral) (auto simp: fds_nth_fds mangoldt_def) finally have sum1: "(\x. if x = Suc 0 then l * g 1 else fds_nth f x * g x * mangoldt x /\<^sub>R ln (real x) / nat_power x s) abs_summable_on Collect primepow" by (rule abs_summable_on_subset) auto also have "?this \ (\x. fds_nth f x * g x * mangoldt x /\<^sub>R ln (real x) / nat_power x s) abs_summable_on Collect primepow" by (intro abs_summable_on_cong) (insert primepow_gt_Suc_0, auto) also have "\ \ (\(p,k). fds_nth f (p ^ Suc k) * g (p ^ Suc k) * mangoldt (p ^ Suc k) /\<^sub>R ln (real (p ^ Suc k)) / nat_power (p ^ Suc k) s) abs_summable_on (?P \ UNIV)" using bij_betw_primepows unfolding case_prod_unfold by (intro abs_summable_on_reindex_bij_betw [symmetric]) also have "\ \ ?th1" by (intro abs_summable_on_cong) (auto simp: f.mult f.power mangoldt_def aprimedivisor_prime_power ln_realpow prime_gt_0_nat nat_power_power_left divide_simps scaleR_conv_of_real simp del: power_Suc) finally show ?th1 . have "eval_fds h s = (\\<^sub>an. fds_nth h n / nat_power n s)" using * unfolding eval_fds_def by (subst infsetsum_nat') auto also have "\ = (\\<^sub>an \ {1} \ {n. primepow n}. if n = 1 then l * g 1 else fds_nth f n * g n * mangoldt n /\<^sub>R ln (real n) / nat_power n s)" by (intro infsetsum_cong_neutral) (auto simp: eq fds_nth_fds mangoldt_def) also have "\ = l * g 1 + (\\<^sub>an | primepow n. if n = 1 then l * g 1 else fds_nth f n * g n * mangoldt n /\<^sub>R ln (real n) / nat_power n s)" (is "_ = _ + ?x") using sum1 primepow_gt_Suc_0 by (subst infsetsum_Un_disjoint) auto also have "?x = (\\<^sub>an\Collect primepow. fds_nth f n * g n * mangoldt n /\<^sub>R ln (real n) / nat_power n s)" (is "_ = infsetsum ?f _") by (intro infsetsum_cong refl) (insert primepow_gt_Suc_0, auto) also have "\ = (\\<^sub>a(p,k)\(?P \ UNIV). fds_nth f (p ^ Suc k) * g (p ^ Suc k) * mangoldt (p ^ Suc k) /\<^sub>R ln (p ^ Suc k) / nat_power (p ^ Suc k) s)" using bij_betw_primepows unfolding case_prod_unfold by (intro infsetsum_reindex_bij_betw [symmetric]) also have "\ = (\\<^sub>a(p,k)\(?P \ UNIV). (fds_nth f p / nat_power p s) ^ Suc k * g (p ^ Suc k) / of_nat (Suc k))" by (intro infsetsum_cong) (auto simp: f.mult f.power mangoldt_def aprimedivisor_prime_power ln_realpow prime_gt_0_nat nat_power_power_left divide_simps scaleR_conv_of_real simp del: power_Suc) finally show ?th2 . qed lemma eval_fds_ln_completely_multiplicative: fixes s :: "'a :: dirichlet_series" and l :: 'a and f :: "'a fds" assumes "completely_multiplicative_function (fds_nth f)" and "fds_nth f 1 \ 0" assumes "s \ 1 > abs_conv_abscissa (fds_deriv f / f)" shows "(\(p,k). (fds_nth f p / nat_power p s) ^ Suc k / of_nat (Suc k)) abs_summable_on ({p. prime p} \ UNIV)" (is ?th1) and "eval_fds (fds_ln l f) s = l + (\\<^sub>a(p, k)\{p. prime p}\UNIV. (fds_nth f p / nat_power p s) ^ Suc k / of_nat (Suc k))" (is ?th2) proof - from assms have "fds_abs_converges (fds_ln l f) s" by (intro fds_abs_converges_ln) (auto intro!: fds_abs_converges_mult intro: fds_abs_converges) hence "fds_abs_converges (fds (\n. fds_nth (fds_ln l f) n * 1)) s" by simp from eval_fds_ln_completely_multiplicative_strong [OF this assms(1,2)] show ?th1 ?th2 by simp_all qed subsection \Exponential and logarithm\ lemma summable_fds_exp_aux: assumes "fds_nth f' 1 = (0 :: 'a :: real_normed_algebra_1)" shows "summable (\k. fds_nth (f' ^ k) n /\<^sub>R fact k)" proof (rule summable_finite) fix k assume "k \ {..n}" hence "n < k" by simp also have "\ < 2 ^ k" by (induction k) auto finally have "fds_nth (f' ^ k) n = 0" using assms by (intro fds_nth_power_eq_0) auto thus "fds_nth (f' ^ k) n /\<^sub>R fact k = 0" by simp qed auto lemma fixes f :: "'a :: dirichlet_series fds" assumes "fds_abs_converges f s" shows fds_abs_converges_exp: "fds_abs_converges (fds_exp f) s" and eval_fds_exp: "eval_fds (fds_exp f) s = exp (eval_fds f s)" proof - have conv: "fds_abs_converges (fds_exp f) s" and ev: "eval_fds (fds_exp f) s = exp (eval_fds f s)" if "fds_abs_converges f s" and [simp]: "fds_nth f (Suc 0) = 0" for f proof - have [simp]: "fds (\n. if n = Suc 0 then 0 else fds_nth f n) = f" by (intro fds_eqI) simp_all have "(\(k,n). fds_nth (f ^ k) n / fact k / nat_power n s) abs_summable_on (UNIV \ {1..})" proof (subst abs_summable_on_Sigma_iff, safe, goal_cases) case (3 k) from that have "fds_abs_converges (f ^ k) s" by (intro fds_abs_converges_power) hence "(\n. fds_nth (f ^ k) n / nat_power n s * inverse (fact k)) abs_summable_on {1..}" unfolding fds_abs_converges_altdef by (intro abs_summable_on_cmult_left) thus ?case by (simp add: field_simps) next case 4 show ?case unfolding abs_summable_on_nat_iff' proof (rule summable_comparison_test_ev[OF always_eventually[OF allI]]) fix k :: nat from that have *: "fds_abs_converges (fds_norm (f ^ k)) (s \ 1)" by (auto simp: fds_abs_converges_power) have "(\\<^sub>an\{1..}. norm (fds_nth (f ^ k) n / fact k / nat_power n s)) = (\\<^sub>an\{1..}. fds_nth (fds_norm (f ^ k)) n / nat_power n (s \ 1) / fact k)" (is "?S = _") by (intro infsetsum_cong) (simp_all add: norm_divide norm_mult norm_nat_power) also have "\ = (\\<^sub>an\{1..}. fds_nth (fds_norm (f ^ k)) n / nat_power n (s \ 1)) /\<^sub>R fact k" (is "_ = ?S' /\<^sub>R _") using * unfolding fds_abs_converges_altdef by (subst infsetsum_cdiv) (auto simp: abs_summable_on_nat_iff scaleR_conv_of_real divide_simps) also have "?S' = eval_fds (fds_norm (f ^ k)) (s \ 1)" using * unfolding fds_abs_converges_altdef eval_fds_def by (subst infsetsum_nat) (auto intro!: suminf_cong) finally have eq: "?S = \ /\<^sub>R fact k" . note eq also have "?S \ 0" by (intro infsetsum_nonneg) auto hence "?S = norm (norm ?S)" by simp also have "eval_fds (fds_norm (f ^ k)) (s \ 1) \ eval_fds (fds_norm f) (s \ 1) ^ k" using that by (intro eval_fds_norm_power_le) auto finally show "norm (norm (\\<^sub>an\{1..}. norm (fds_nth (f ^ k) n / fact k / nat_power n s))) \ eval_fds (fds_norm f) (s \ 1) ^ k /\<^sub>R fact k" by (simp add: divide_right_mono) next from exp_converges[of "eval_fds (fds_norm f) (s \ 1)"] show "summable (\x. eval_fds (fds_norm f) (s \ 1) ^ x /\<^sub>R fact x)" by (simp add: sums_iff) qed qed auto hence summable: "(\(n,k). fds_nth (f ^ k) n / fact k / nat_power n s) abs_summable_on {1..} \ UNIV" by (subst abs_summable_on_Times_swap) (simp add: case_prod_unfold) have summable': "(\k. fds_nth (f ^ k) n / fact k) abs_summable_on UNIV" for n using abs_summable_on_cmult_left[of "nat_power n s", OF abs_summable_on_Sigma_project2 [OF summable, of n]] by (cases "n = 0") simp_all have "(\n. \\<^sub>ak. fds_nth (f ^ k) n / fact k / nat_power n s) abs_summable_on {1..}" using summable by (rule abs_summable_on_Sigma_project1') auto also have "?this \ (\n. (\k. fds_nth (f ^ k) n / fact k) * inverse (nat_power n s)) abs_summable_on {1..}" proof (intro abs_summable_on_cong refl, goal_cases) case (1 n) hence "(\\<^sub>ak. fds_nth (f ^ k) n / fact k / nat_power n s) = (\\<^sub>ak. fds_nth (f ^ k) n / fact k) * inverse (nat_power n s)" using summable'[of n] by (subst infsetsum_cmult_left [symmetric]) (auto simp: field_simps) also have "(\\<^sub>ak. fds_nth (f ^ k) n / fact k) = (\k. fds_nth (f ^ k) n / fact k)" using summable'[of n] 1 by (intro abs_summable_on_cong refl infsetsum_nat') auto finally show ?case . qed finally show "fds_abs_converges (fds_exp f) s" by (simp add: fds_exp_def fds_nth_fds' abs_summable_on_Sigma_iff scaleR_conv_of_real fds_abs_converges_altdef field_simps) have "eval_fds (fds_exp f) s = (\n. (\k. fds_nth (f ^ k) n /\<^sub>R fact k) / nat_power n s)" by (simp add: fds_exp_def eval_fds_def fds_nth_fds') also have "\ = (\n. (\\<^sub>ak. fds_nth (f ^ k) n /\<^sub>R fact k) / nat_power n s)" proof (intro suminf_cong, goal_cases) case (1 n) show ?case proof (cases "n = 0") case False have "(\k. fds_nth (f ^ k) n /\<^sub>R fact k) = (\\<^sub>ak. fds_nth (f ^ k) n /\<^sub>R fact k)" using summable'[of n] False by (intro infsetsum_nat' [symmetric]) (auto simp: scaleR_conv_of_real field_simps) thus ?thesis by simp qed simp_all qed also have "\ = (\\<^sub>an. (\\<^sub>ak. fds_nth (f ^ k) n /\<^sub>R fact k) / nat_power n s)" proof (intro infsetsum_nat' [symmetric], goal_cases) case 1 have *: "UNIV - {Suc 0..} = {0}" by auto have "(\x. \\<^sub>ay. fds_nth (f ^ y) x / fact y / nat_power x s) abs_summable_on {1..}" by (intro abs_summable_on_Sigma_project1'[OF summable]) auto also have "?this \ (\x. (\\<^sub>ay. fds_nth (f ^ y) x / fact y) * inverse (nat_power x s)) abs_summable_on {1..}" using summable' by (intro abs_summable_on_cong refl, subst infsetsum_cmult_left [symmetric]) (auto simp: field_simps) also have "\ \ (\x. (\\<^sub>ay. fds_nth (f ^ y) x /\<^sub>R fact y) / (nat_power x s)) abs_summable_on {1..}" by (simp add: field_simps scaleR_conv_of_real) finally show ?case by (rule abs_summable_on_finite_diff) (use * in auto) qed also have "\ = (\\<^sub>an. (\\<^sub>ak. fds_nth (f ^ k) n /\<^sub>R fact k * inverse (nat_power n s)))" using summable' by (subst infsetsum_cmult_left) (auto simp: field_simps scaleR_conv_of_real) also have "\ = (\\<^sub>an\{1..}. (\\<^sub>ak. fds_nth (f ^ k) n /\<^sub>R fact k * inverse (nat_power n s)))" by (intro infsetsum_cong_neutral) (auto simp: Suc_le_eq) also have "\ = (\\<^sub>ak. \\<^sub>an\{1..}. fds_nth (f ^ k) n / nat_power n s /\<^sub>R fact k)" using summable by (subst infsetsum_swap) (auto simp: field_simps scaleR_conv_of_real case_prod_unfold) also have "\ = (\\<^sub>ak. (\\<^sub>an\{1..}. fds_nth (f ^ k) n / nat_power n s) /\<^sub>R fact k)" by (subst infsetsum_scaleR_right) simp also have "\ = (\\<^sub>ak. eval_fds f s ^ k /\<^sub>R fact k)" proof (intro infsetsum_cong refl, goal_cases) case (1 k) have *: "fds_abs_converges (f ^ k) s" by (intro fds_abs_converges_power that) have "(\\<^sub>an\{1..}. fds_nth (f ^ k) n / nat_power n s) = (\\<^sub>an. fds_nth (f ^ k) n / nat_power n s)" by (intro infsetsum_cong_neutral) (auto simp: Suc_le_eq) also have "\ = eval_fds (f ^ k) s" using * unfolding eval_fds_def by (intro infsetsum_nat') (auto simp: fds_abs_converges_def abs_summable_on_nat_iff') also from that have "\ = eval_fds f s ^ k" by (simp add: eval_fds_power) finally show ?case by simp qed also have "\ = (\k. eval_fds f s ^ k /\<^sub>R fact k)" using exp_converges[of "norm (eval_fds f s)"] by (intro infsetsum_nat') (auto simp: abs_summable_on_nat_iff' sums_iff field_simps norm_power) also have "\ = exp (eval_fds f s)" by (simp add: exp_def) finally show "eval_fds (fds_exp f) s = exp (eval_fds f s)" . qed define f' where "f' = f - fds_const (fds_nth f 1)" have *: "fds_abs_converges (fds_exp f') s" by (auto simp: f'_def intro!: fds_abs_converges_diff conv assms) have "fds_abs_converges (fds_const (exp (fds_nth f 1)) * fds_exp f') s" unfolding f'_def by (intro fds_abs_converges_mult conv fds_abs_converges_diff assms) auto thus "fds_abs_converges (fds_exp f) s" unfolding f'_def by (simp add: fds_exp_times_fds_nth_0) have "eval_fds (fds_exp f) s = eval_fds (fds_const (exp (fds_nth f 1)) * fds_exp f') s" by (simp add: f'_def fds_exp_times_fds_nth_0) also have "\ = exp (fds_nth f (Suc 0)) * eval_fds (fds_exp f') s" using * using assms by (subst eval_fds_mult) (simp_all) also have "\ = exp (eval_fds f s)" using ev[of f'] assms unfolding f'_def by (auto simp: fds_abs_converges_diff eval_fds_diff fds_abs_converges_imp_converges exp_diff) finally show "eval_fds (fds_exp f) s = exp (eval_fds f s)" . qed lemma fds_exp_add: fixes f :: "'a :: dirichlet_series fds" shows "fds_exp (f + g) = fds_exp f * fds_exp g" proof (rule fds_eqI_truncate) fix m :: nat assume m: "m > 0" let ?T = "fds_truncate m" have "?T (fds_exp (f + g)) = ?T (fds_exp (?T f + ?T g))" by (simp add: fds_truncate_exp fds_truncate_add_strong [symmetric]) also have "fds_exp (?T f + ?T g) = fds_exp (?T f) * fds_exp (?T g)" proof (rule eval_fds_eqD) have "fds_abs_converges (fds_exp (?T f + ?T g)) 0" by (intro fds_abs_converges_exp fds_abs_converges_add) auto thus "conv_abscissa (fds_exp (?T f + ?T g)) < \" using conv_abscissa_PInf_iff by blast hence "fds_abs_converges (fds_exp (fds_truncate m f) * fds_exp (fds_truncate m g)) 0" by (intro fds_abs_converges_mult fds_abs_converges_exp) auto thus "conv_abscissa (fds_exp (fds_truncate m f) * fds_exp (fds_truncate m g)) < \" using conv_abscissa_PInf_iff by blast show "frequently (\s. eval_fds (fds_exp (fds_truncate m f + fds_truncate m g)) s = eval_fds (fds_exp (fds_truncate m f) * fds_exp (fds_truncate m g)) s) ((\s. s \ 1) going_to at_top)" by (auto simp: eval_fds_add eval_fds_mult eval_fds_exp fds_abs_converges_add fds_abs_converges_exp exp_add) qed also have "?T \ = ?T (fds_exp f * fds_exp g)" by (subst fds_truncate_mult [symmetric], subst (1 2) fds_truncate_exp) (simp add: fds_truncate_mult) finally show "?T (fds_exp (f + g)) = \" . qed lemma fds_exp_minus: fixes f :: "'a :: dirichlet_series fds" shows "fds_exp (-f) = inverse (fds_exp f)" proof (rule fds_right_inverse_unique) have "fds_exp f * fds_exp (- f) = fds_exp (f + (-f))" by (subst fds_exp_add) simp_all also have "f + (-f) = 0" by simp also have "fds_exp \ = 1" by simp finally show "fds_exp f * fds_exp (-f) = 1" . qed lemma abs_conv_abscissa_exp: fixes f :: "'a :: dirichlet_series fds" shows "abs_conv_abscissa (fds_exp f) \ abs_conv_abscissa f" by (intro abs_conv_abscissa_mono fds_abs_converges_exp) lemma fds_deriv_exp [simp]: fixes f :: "'a :: dirichlet_series fds" shows "fds_deriv (fds_exp f) = fds_exp f * fds_deriv f" proof (rule fds_eqI_truncate) fix m :: nat assume m: "m > 0" let ?T = "fds_truncate m" have "abs_conv_abscissa (fds_deriv (?T f)) = -\" by (simp add: abs_conv_abscissa_deriv) have "?T (fds_deriv (fds_exp f)) = ?T (fds_deriv (fds_exp (?T f)))" by (simp add: fds_truncate_deriv fds_truncate_exp) also have "fds_deriv (fds_exp (?T f)) = fds_exp (?T f) * fds_deriv (?T f)" proof (rule eval_fds_eqD) note abscissa = conv_le_abs_conv_abscissa abs_conv_abscissa_exp note abscissa' = abscissa[THEN le_less_trans] have "fds_abs_converges (fds_deriv (fds_exp (fds_truncate m f))) 0" by (intro fds_abs_converges ) (auto simp: abs_conv_abscissa_deriv intro: le_less_trans[OF abs_conv_abscissa_exp]) thus "conv_abscissa (fds_deriv (fds_exp (fds_truncate m f))) < \" using conv_abscissa_PInf_iff by blast have "fds_abs_converges (fds_exp (fds_truncate m f) * fds_deriv (fds_truncate m f)) 0" by (intro fds_abs_converges_mult fds_abs_converges_exp) (auto intro: fds_abs_converges simp add: fds_truncate_deriv [symmetric]) thus "conv_abscissa (fds_exp (fds_truncate m f) * fds_deriv (fds_truncate m f)) < \" using conv_abscissa_PInf_iff by blast show "\\<^sub>F s in (\s. s \ 1) going_to at_top. eval_fds (fds_deriv (fds_exp (?T f))) s = eval_fds (fds_exp (?T f) * fds_deriv (?T f)) s" proof (intro always_eventually eventually_frequently allI, goal_cases) case (2 s) have "eval_fds (fds_deriv (fds_exp (?T f))) s = deriv (eval_fds (fds_exp (?T f))) s" by (auto simp: eval_fds_exp eval_fds_mult fds_abs_converges_mult fds_abs_converges_exp fds_abs_converges eval_fds_deriv abscissa') also have "eval_fds (fds_exp (?T f)) = (\s. exp (eval_fds (?T f) s))" by (intro ext eval_fds_exp) auto also have "deriv \ = (\s. exp (eval_fds (?T f) s) * deriv (eval_fds (?T f)) s)" by (auto intro!: DERIV_imp_deriv derivative_eq_intros simp: eval_fds_deriv) also have "\ = eval_fds (fds_exp (?T f) * fds_deriv (?T f))" by (auto simp: eval_fds_exp eval_fds_mult fds_abs_converges_mult fds_abs_converges_exp fds_abs_converges eval_fds_deriv abs_conv_abscissa_deriv) finally show ?case . qed auto qed also have "?T \ = ?T (fds_exp f * fds_deriv f)" by (subst fds_truncate_mult [symmetric]) (simp add: fds_truncate_exp fds_truncate_deriv [symmetric], simp add: fds_truncate_mult) finally show "?T (fds_deriv (fds_exp f)) = \" . qed lemma fds_exp_ln_strong: fixes f :: "'a :: dirichlet_series fds" assumes "fds_nth f (Suc 0) \ 0" shows "fds_exp (fds_ln l f) = fds_const (exp l / fds_nth f (Suc 0)) * f" proof - let ?c = "exp l / fds_nth f (Suc 0)" have "f * fds_const ?c = f * (fds_exp (-fds_ln l f) * fds_exp (fds_ln l f)) * fds_const ?c" (is "_ = _ * (?g * ?h) * _") by (subst fds_exp_add [symmetric]) simp also have "\ = fds_const ?c * (f * ?g) * ?h" by (simp add: mult_ac) also have "f * ?g = fds_const (inverse ?c)" proof (rule fds_deriv_eq_imp_eq) have "fds_deriv (f * fds_exp (-fds_ln l f)) = fds_exp (- fds_ln l f) * fds_deriv f * (1 - f / f)" by (simp add: divide_fds_def algebra_simps) also from assms have "f / f = 1" by (simp add: divide_fds_def fds_right_inverse) finally show "fds_deriv (f * fds_exp (-fds_ln l f)) = fds_deriv (fds_const (inverse ?c))" by simp qed (insert assms, auto simp: exp_minus field_simps) also have "fds_const ?c * fds_const (inverse ?c) = 1" using assms by (subst fds_const_mult [symmetric]) (simp add: divide_simps) finally show ?thesis by (simp add: mult_ac) qed lemma fds_exp_ln [simp]: fixes f :: "'a :: dirichlet_series fds" assumes "exp l = fds_nth f (Suc 0)" shows "fds_exp (fds_ln l f) = f" using assms by (subst fds_exp_ln_strong) auto lemma fds_ln_exp [simp]: fixes f :: "'a :: dirichlet_series fds" assumes "l = fds_nth f (Suc 0)" shows "fds_ln l (fds_exp f) = f" proof (rule fds_deriv_eq_imp_eq) have "fds_deriv (fds_ln l (fds_exp f)) = fds_deriv f * (fds_exp f / fds_exp f)" by (simp add: algebra_simps divide_fds_def) also have "fds_exp f / fds_exp f = 1" by (simp add: divide_fds_def fds_right_inverse) finally show "fds_deriv (fds_ln l (fds_exp f)) = fds_deriv f" by simp qed (insert assms, auto simp: field_simps) subsection \Euler products\ lemma fds_euler_product_LIMSEQ: fixes f :: "'a :: {nat_power, real_normed_field, banach, second_countable_topology} fds" assumes "multiplicative_function (fds_nth f)" and "fds_abs_converges f s" shows "(\n. \p\n. if prime p then \i. fds_nth f (p ^ i) / nat_power (p ^ i) s else 1) \ eval_fds f s" unfolding eval_fds_def proof (rule euler_product_LIMSEQ) show "multiplicative_function (\n. fds_nth f n / nat_power n s)" by (rule multiplicative_function_divide_nat_power) fact+ qed (insert assms, auto simp: fds_abs_converges_def) lemma fds_euler_product_LIMSEQ': fixes f :: "'a :: {nat_power, real_normed_field, banach, second_countable_topology} fds" assumes "completely_multiplicative_function (fds_nth f)" and "fds_abs_converges f s" shows "(\n. \p\n. if prime p then inverse (1 - fds_nth f p / nat_power p s) else 1) \ eval_fds f s" unfolding eval_fds_def proof (rule euler_product_LIMSEQ') show "completely_multiplicative_function (\n. fds_nth f n / nat_power n s)" by (rule completely_multiplicative_function_divide_nat_power) fact+ qed (insert assms, auto simp: fds_abs_converges_def) lemma fds_abs_convergent_euler_product: fixes f :: "'a :: {nat_power, real_normed_field, banach, second_countable_topology} fds" assumes "multiplicative_function (fds_nth f)" and "fds_abs_converges f s" shows "abs_convergent_prod (\p. if prime p then \i. fds_nth f (p ^ i) / nat_power (p ^ i) s else 1)" unfolding eval_fds_def proof (rule abs_convergent_euler_product) show "multiplicative_function (\n. fds_nth f n / nat_power n s)" by (rule multiplicative_function_divide_nat_power) fact+ qed (insert assms, auto simp: fds_abs_converges_def) lemma fds_abs_convergent_euler_product': fixes f :: "'a :: {nat_power, real_normed_field, banach, second_countable_topology} fds" assumes "completely_multiplicative_function (fds_nth f)" and "fds_abs_converges f s" shows "abs_convergent_prod (\p. if prime p then inverse (1 - fds_nth f p / nat_power p s) else 1)" unfolding eval_fds_def proof (rule abs_convergent_euler_product') show "completely_multiplicative_function (\n. fds_nth f n / nat_power n s)" by (rule completely_multiplicative_function_divide_nat_power) fact+ qed (insert assms, auto simp: fds_abs_converges_def) lemma fds_abs_convergent_zero_iff: fixes f :: "'a :: {nat_power_field, real_normed_field, banach, second_countable_topology} fds" assumes "completely_multiplicative_function (fds_nth f)" assumes "fds_abs_converges f s" shows "eval_fds f s = 0 \ (\p. prime p \ fds_nth f p = nat_power p s)" proof - let ?g = "\p. if prime p then inverse (1 - fds_nth f p / nat_power p s) else 1" have lim: "(\n. \p\n. ?g p) \ eval_fds f s" by (intro fds_euler_product_LIMSEQ' assms) have conv: "convergent_prod ?g" by (intro abs_convergent_prod_imp_convergent_prod fds_abs_convergent_euler_product' assms) { assume "eval_fds f s = 0" from convergent_prod_to_zero_iff[OF conv] and this and lim have "\p. prime p \ fds_nth f p = nat_power p s" by (auto split: if_splits) } moreover { assume "\p. prime p \ fds_nth f p = nat_power p s" then obtain p where "prime p" "fds_nth f p = nat_power p s" by blast moreover from this have "nat_power p s \ 0" by (intro nat_power_nonzero) (auto simp: prime_gt_0_nat) ultimately have "(\n. \p\n. ?g p) \ 0" using convergent_prod_to_zero_iff[OF conv] by (auto intro!: exI[of _ p] split: if_splits) from tendsto_unique[OF _ lim this] have "eval_fds f s = 0" by simp } ultimately show ?thesis by blast qed lemma fixes s :: "'a :: {nat_power_normed_field,banach,euclidean_space}" assumes "s \ 1 > 1" shows euler_product_fds_zeta: "(\n. \p\n. if prime p then inverse (1 - 1 / nat_power p s) else 1) \ eval_fds fds_zeta s" (is ?th1) and eval_fds_zeta_nonzero: "eval_fds fds_zeta s \ 0" proof - have *: "completely_multiplicative_function (fds_nth fds_zeta)" by standard auto have lim: "(\n. \p\n. if prime p then inverse (1 - fds_nth fds_zeta p / nat_power p s) else 1) \ eval_fds fds_zeta s" (is "filterlim ?g _ _") using assms by (intro fds_euler_product_LIMSEQ' * fds_abs_summable_zeta) also have "?g = (\n. \p\n. if prime p then inverse (1 - 1 / nat_power p s) else 1)" by (intro ext prod.cong refl) (auto simp: fds_zeta_def fds_nth_fds) finally show ?th1 . { fix p :: nat assume "prime p" from this have "p > 1" by (simp add: prime_gt_Suc_0_nat) hence "norm (nat_power p s) = real p powr (s \ 1)" by (simp add: norm_nat_power) also have "\ > real p powr 0" using assms and \p > 1\ by (intro powr_less_mono) auto finally have "nat_power p s \ 1" using \p > 1\ by auto } hence **: "\p. prime p \ fds_nth fds_zeta p = nat_power p s" by (auto simp: fds_zeta_def fds_nth_fds) show "eval_fds fds_zeta s \ 0" using assms * ** by (subst fds_abs_convergent_zero_iff) simp_all qed lemma fds_primepow_subseries_euler_product_cm: fixes f :: "'a :: dirichlet_series fds" assumes "completely_multiplicative_function (fds_nth f)" "prime p" assumes "s \ 1 > abs_conv_abscissa f" shows "eval_fds (fds_primepow_subseries p f) s = 1 / (1 - fds_nth f p / nat_power p s)" proof - let ?f = "(\n. \pa\n. if prime pa then inverse (1 - fds_nth (fds_primepow_subseries p f) pa / nat_power pa s) else 1)" have "sequentially \ bot" by simp moreover have "?f \ eval_fds (fds_primepow_subseries p f) s" by (intro fds_euler_product_LIMSEQ' completely_multiplicative_function_only_pows assms fds_abs_converges_subseries) (insert assms, auto intro!: fds_abs_converges) moreover have "eventually (\n. ?f n = 1 / (1 - fds_nth f p / nat_power p s)) at_top" using eventually_ge_at_top[of p] proof eventually_elim case (elim n) have "(\pa\n. if prime pa then inverse (1 - fds_nth (fds_primepow_subseries p f) pa / nat_power pa s) else 1) = (\q\n. if q = p then inverse (1 - fds_nth f p / nat_power p s) else 1)" using \prime p\ by (intro prod.cong) (auto simp: fds_nth_subseries prime_prime_factors) also have "\ = 1 / (1 - fds_nth f p / nat_power p s)" using elim by (subst prod.delta) (auto simp: divide_simps) finally show ?case . qed hence "?f \ 1 / (1 - fds_nth f p / nat_power p s)" by (rule tendsto_eventually) ultimately show ?thesis by (rule tendsto_unique) qed subsection \Non-negative Dirichlet series\ lemma nonneg_Reals_sum: "(\x. x \ A \ f x \ \\<^sub>\\<^sub>0) \ sum f A \ \\<^sub>\\<^sub>0" by (induction A rule: infinite_finite_induct) auto locale nonneg_dirichlet_series = fixes f :: "'a :: dirichlet_series fds" assumes nonneg_coeffs_aux: "n > 0 \ fds_nth f n \ \\<^sub>\\<^sub>0" begin lemma nonneg_coeffs: "fds_nth f n \ \\<^sub>\\<^sub>0" using nonneg_coeffs_aux[of n] by (cases "n = 0") auto end lemma nonneg_dirichlet_series_0 [simp,intro]: "nonneg_dirichlet_series 0" by standard (auto simp: zero_fds_def) lemma nonneg_dirichlet_series_1 [simp,intro]: "nonneg_dirichlet_series 1" by standard (auto simp: one_fds_def) lemma nonneg_dirichlet_series_const [simp,intro]: "c \ \\<^sub>\\<^sub>0 \ nonneg_dirichlet_series (fds_const c)" by standard (auto simp: fds_const_def) lemma nonneg_dirichlet_series_add [intro]: assumes "nonneg_dirichlet_series f" "nonneg_dirichlet_series g" shows "nonneg_dirichlet_series (f + g)" proof - interpret f: nonneg_dirichlet_series f by fact interpret g: nonneg_dirichlet_series g by fact show ?thesis by standard (auto intro!: nonneg_Reals_add_I f.nonneg_coeffs g.nonneg_coeffs) qed lemma nonneg_dirichlet_series_mult [intro]: assumes "nonneg_dirichlet_series f" "nonneg_dirichlet_series g" shows "nonneg_dirichlet_series (f * g)" proof - interpret f: nonneg_dirichlet_series f by fact interpret g: nonneg_dirichlet_series g by fact show ?thesis by standard (auto intro!: nonneg_Reals_sum nonneg_Reals_mult_I f.nonneg_coeffs g.nonneg_coeffs simp: fds_nth_mult dirichlet_prod_def) qed lemma nonneg_dirichlet_series_power [intro]: assumes "nonneg_dirichlet_series f" shows "nonneg_dirichlet_series (f ^ n)" using assms by (induction n) auto context nonneg_dirichlet_series begin lemma nonneg_exp [intro]: "nonneg_dirichlet_series (fds_exp f)" proof fix n :: nat assume "n > 0" define c where "c = exp (fds_nth f (Suc 0))" define f' where "f' = fds (\n. if n = Suc 0 then 0 else fds_nth f n)" from nonneg_coeffs[of 1] obtain c' where "fds_nth f (Suc 0) = of_real c'" by (auto elim!: nonneg_Reals_cases) hence "c = of_real (exp c')" by (simp add: c_def exp_of_real) hence c: "c \ \\<^sub>\\<^sub>0" by simp have less: "n < 2 ^ k" if "n < k" for k proof - have "n < k" by fact also have "\ < 2 ^ k" by (induction k) auto finally show ?thesis . qed have nonneg_power: "fds_nth (f' ^ k) n \ \\<^sub>\\<^sub>0" for k proof - have "nonneg_dirichlet_series f'" by standard (insert nonneg_coeffs, auto simp: f'_def) interpret nonneg_dirichlet_series "f' ^ k" by (intro nonneg_dirichlet_series_power) fact+ from nonneg_coeffs[of n] show ?thesis . qed hence "fds_nth (fds_exp f) n = c * (\k. fds_nth (f' ^ k) n /\<^sub>R fact k)" by (simp add: fds_exp_def fds_nth_fds' f'_def c_def) also have "(\k. fds_nth (f' ^ k) n /\<^sub>R fact k) = (\k\n. fds_nth (f' ^ k) n /\<^sub>R fact k)" by (intro suminf_finite) (auto intro!: fds_nth_power_eq_0 less simp: f'_def not_le) also have "c * \ \ \\<^sub>\\<^sub>0" unfolding scaleR_conv_of_real by (intro nonneg_Reals_mult_I nonneg_Reals_sum nonneg_power, unfold nonneg_Reals_of_real_iff ) (auto simp: c) finally show "fds_nth (fds_exp f) n \ \\<^sub>\\<^sub>0" . qed end lemma nonneg_dirichlet_series_lnD: assumes "nonneg_dirichlet_series (fds_ln l f)" "exp l = fds_nth f (Suc 0)" shows "nonneg_dirichlet_series f" proof - from assms have "nonneg_dirichlet_series (fds_exp (fds_ln l f))" by (intro nonneg_dirichlet_series.nonneg_exp) thus ?thesis using assms by simp qed context nonneg_dirichlet_series begin lemma fds_of_real_norm: "fds_of_real (fds_norm f) = f" proof (rule fds_eqI) fix n :: nat assume n: "n > 0" show "fds_nth (fds_of_real (fds_norm f)) n = fds_nth f n" using nonneg_coeffs[of n] by (auto elim!: nonneg_Reals_cases) qed end (* TODO: Simplify proof *) lemma pringsheim_landau_aux: fixes c :: real and f :: "complex fds" assumes "nonneg_dirichlet_series f" assumes abscissa: "c \ abs_conv_abscissa f" assumes g: "\s. s \ A \ Re s > c \ g s = eval_fds f s" assumes "g holomorphic_on A" "open A" "c \ A" shows "\x. x < c \ fds_abs_converges f (of_real x)" proof - interpret nonneg_dirichlet_series f by fact define a where "a = 1 + c" define g' where "g' = (\s. if s \ {s. Re s > c} then eval_fds f s else g s)" \ \We can find some $\varepsilon > 0$ such that the Dirichlet series can be continued analytically in a ball of radius $1 + \varepsilon$ around $a$.\ from \open A\ \c \ A\ obtain \ where \: "\ > 0" "ball c \ \ A" by (auto simp: open_contains_ball) define \ where "\ = sqrt (1 + \^2) - 1" from \ have \: "\ > 0" by (simp add: \_def) have ball_a_subset: "ball a (1 + \) \ {s. Re s > c} \ A" proof (intro subsetI) fix s :: complex assume s: "s \ ball a (1 + \)" define x y where "x = Re s" and "y = Im s" have [simp]: "s = x + \ * y" by (simp add: complex_eq_iff x_def y_def) show "s \ {s. Re s > c} \ A" proof (cases "Re s \ c") case True hence "(c - x)\<^sup>2 + y\<^sup>2 \ (1 + c - x)\<^sup>2 + y\<^sup>2 - 1" by (simp add: power2_eq_square algebra_simps) also from s have "(1 + c - x)\<^sup>2 + y\<^sup>2 - 1 < \\<^sup>2" by (auto simp: dist_norm cmod_def a_def \_def) finally have "sqrt ((c - x)\<^sup>2 + y\<^sup>2) < \" using \ by (intro real_less_lsqrt) auto hence "s \ ball c \" by (auto simp: dist_norm cmod_def) also have "\ \ A" by fact finally show ?thesis .. qed auto qed have holo: "g' holomorphic_on ball a (1 + \)" unfolding g'_def proof (intro holomorphic_on_subset[OF _ ball_a_subset] holomorphic_on_If_Un) have "conv_abscissa f \ abs_conv_abscissa f" by (rule conv_le_abs_conv_abscissa) also have "\ \ ereal c" by fact finally have*: "conv_abscissa f \ ereal c" . show "eval_fds f holomorphic_on {s. c < Re s}" by (intro holomorphic_intros) (auto intro: le_less_trans[OF *]) qed (insert assms, auto intro!: holomorphic_intros open_halfspace_Re_gt) define f' where "f' = fds_norm f" have f_f': "f = fds_of_real f'" by (simp add: f'_def fds_of_real_norm) have f'_nonneg: "fds_nth f' n \ 0" for n using nonneg_coeffs[of n] by (auto elim!: nonneg_Reals_cases simp: f'_def) have deriv: "(\n. (deriv ^^ n) g' a) = (\n. eval_fds ((fds_deriv ^^ n) f) a)" proof fix n :: nat have ev: "eventually (\s. s \ {s. Re s > c}) (nhds (complex_of_real a))" by (intro eventually_nhds_in_open open_halfspace_Re_gt) (auto simp: a_def) have "(deriv ^^ n) g' a = (deriv ^^ n) (eval_fds f) a" by (intro higher_deriv_cong_ev refl eventually_mono[OF ev]) (auto simp: g'_def) also have "\ = eval_fds ((fds_deriv ^^ n) f) a" proof (intro eval_fds_higher_deriv [symmetric]) have "conv_abscissa f \ abs_conv_abscissa f" by (rule conv_le_abs_conv_abscissa) also have "\ \ ereal c" by (rule assms) also have "\ < a" by (simp add: a_def) finally show "conv_abscissa f < ereal (complex_of_real a \ 1)" by simp qed finally show "(deriv ^^ n) g' a = eval_fds ((fds_deriv ^^ n) f) a" . qed have nth_deriv_conv: "fds_abs_converges ((fds_deriv ^^ n) f) (of_real a)" for n by (intro fds_abs_converges) (auto simp: abs_conv_abscissa_higher_deriv a_def intro!: le_less_trans[OF abscissa]) have nth_deriv_eq: "(fds_deriv ^^ n) f = fds (\k. (-1) ^ n * fds_nth f k * ln (real k) ^ n)" for n proof - have "fds_nth ((fds_deriv ^^ n) f) k = (-1) ^ n * fds_nth f k * ln (real k) ^ n" for k by (induction n) (simp_all add: fds_deriv_def fds_eq_iff fds_nth_fds' scaleR_conv_of_real) thus ?thesis by (intro fds_eqI) simp_all qed have deriv': "(\n. eval_fds ((fds_deriv ^^ n) f) (complex_of_real a)) = (\n. (- 1) ^ n * complex_of_real (\\<^sub>ak. fds_nth f' k * ln (real k) ^ n / real k powr a))" proof fix n have "eval_fds ((fds_deriv ^^ n) f) (of_real a) = (\\<^sub>ak. fds_nth ((fds_deriv ^^ n) f) k / of_nat k powr complex_of_real a)" using nth_deriv_conv by (subst eval_fds_altdef) auto hence "eval_fds ((fds_deriv ^^ n) f) (of_real a) = (\\<^sub>ak. (- 1) ^ n *\<^sub>R (fds_nth f k * ln (real k) ^ n / k powr a))" by (simp add: nth_deriv_eq fds_nth_fds' powr_Reals_eq scaleR_conv_of_real algebra_simps) also have "\ = (- 1) ^ n * (\\<^sub>ak. of_real (fds_nth f' k * ln (real k) ^ n / k powr a))" by (subst infsetsum_scaleR_right) (simp_all add: scaleR_conv_of_real f_f') also have "\ = (- 1) ^ n * of_real (\\<^sub>ak. fds_nth f' k * ln (real k) ^ n / k powr a)" by (subst infsetsum_of_real) (rule refl) finally show "eval_fds ((fds_deriv ^^ n) f) (complex_of_real a) = (- 1) ^ n * complex_of_real (\\<^sub>ak. fds_nth f' k * ln (real k) ^ n / real k powr a)" . qed define s :: complex where "s = c - \ / 2" have s: "Re s < c" using assms \ by (simp_all add: s_def \_def field_simps) have "s \ ball a (1 + \)" using s by (simp add: a_def dist_norm cmod_def s_def) from holomorphic_power_series[OF holo this] have sums: "(\n. (deriv ^^ n) g' a / fact n * (s - a) ^ n) sums g' s" by simp also note deriv also have "s - a = -of_real (1 + \ / 2)" by (simp add: s_def a_def) also have "(\n. \ ^ n) = (\n. of_real ((-1) ^ n * (1 + \ / 2) ^ n))" by (intro ext) (subst power_minus, auto) also have "(\n. eval_fds ((fds_deriv ^^ n) f) a / fact n * \ n) = (\n. of_real ((-1) ^ n * eval_fds ((fds_deriv ^^ n) f') a / fact n * (1+\/2) ^ n))" using nth_deriv_conv by (simp add: f_f' fds_abs_converges_imp_converges mult_ac) finally have "summable \" by (simp add: sums_iff) hence summable: "summable (\n. (-1)^n * eval_fds ((fds_deriv ^^ n) f') a / fact n * (1+\/2)^n)" by (subst (asm) summable_of_real_iff) have "(\(n,k). (-1)^n * fds_nth f k * ln (real k) ^ n / (real k powr a) * ((s-a) ^ n / fact n)) abs_summable_on (UNIV \ UNIV)" proof (subst abs_summable_on_Sigma_iff, safe, goal_cases) case (3 n) from nth_deriv_conv[of n] show ?case unfolding fds_abs_converges_altdef' by (intro abs_summable_on_cmult_left) (simp add: nth_deriv_eq fds_nth_fds' powr_Reals_eq) next case 4 have nth_deriv_f_f': "(fds_deriv ^^ n) f = fds_of_real ((fds_deriv ^^ n) f')" for n by (induction n) (auto simp: f'_def fds_of_real_norm) have norm_nth_deriv_f: "norm (fds_nth ((fds_deriv ^^ n) f) k) = (-1) ^ n * of_real (fds_nth ((fds_deriv ^^ n) f') k)" for n k proof (induction n) case (Suc n) thus ?case by (cases k) (auto simp: f_f' fds_nth_deriv scaleR_conv_of_real norm_mult) qed (auto simp: f'_nonneg f_f') note summable also have "(\n. (-1)^n * eval_fds ((fds_deriv ^^ n) f') a / fact n * (1+\/2)^n) = (\n. \\<^sub>ak. norm ((- 1) ^ n * fds_nth f k * ln (real k) ^ n / (real k powr a) * ((s - a) ^ n / fact n)))" (is "_ = ?h") proof (rule ext, goal_cases) case (1 n) have "(\\<^sub>ak. norm ((- 1) ^ n * fds_nth f k * ln (real k) ^ n / (real k powr a) * ((s - a) ^ n / fact n))) = (norm ((s - a) ^ n / fact n) * (-1) ^ n) *\<^sub>R (\\<^sub>ak. (-1) ^ n * norm (fds_nth ((fds_deriv ^^ n) f) k / real k powr a))" (is "_ = _ *\<^sub>R ?S") by (subst infsetsum_scaleR_right [symmetric]) (auto simp: norm_mult norm_divide norm_power mult_ac nth_deriv_eq fds_nth_fds') also have "?S = (\\<^sub>ak. fds_nth ((fds_deriv ^^ n) f') k / real k powr a)" by (intro infsetsum_cong) (auto simp: norm_mult norm_divide norm_power norm_nth_deriv_f) also have "\ = eval_fds ((fds_deriv ^^ n) f') a" using nth_deriv_conv[of n] by (subst eval_fds_altdef) (auto simp: f'_def nth_deriv_f_f') also have "(norm ((s - a) ^ n / fact n) * (- 1) ^ n) *\<^sub>R eval_fds ((fds_deriv ^^ n) f') a = (-1) ^ n * eval_fds ((fds_deriv ^^ n) f') a / fact n * norm (s - a) ^ n" by (simp add: norm_divide norm_power) also have s_a: "s - a = -of_real (1 + \ / 2)" by (simp add: s_def a_def) have "norm (s - a) = 1 + \ / 2" unfolding s_a norm_minus_cancel norm_of_real using \ by simp finally show ?case .. qed also have "?h n \ 0" for n by (intro infsetsum_nonneg) auto hence "?h = (\n. norm (?h n))" by simp finally show ?case unfolding abs_summable_on_nat_iff' . qed auto hence "(\(k,n). (-1)^n * fds_nth f k * ln (real k) ^ n / (real k powr a) * ((s-a) ^ n / fact n)) abs_summable_on (UNIV \ UNIV)" by (subst (asm) abs_summable_on_Times_swap) (simp add: case_prod_unfold) hence "(\k. \\<^sub>an. (- 1) ^ n * fds_nth f k * ln (real k) ^ n / (k powr a) * ((s - a) ^ n / fact n)) abs_summable_on UNIV" (is "?h abs_summable_on _") by (rule abs_summable_on_Sigma_project1') auto also have "?this \ (\k. fds_nth f k / nat_power k s) abs_summable_on UNIV" proof (intro abs_summable_on_cong refl, goal_cases) case (1 k) have "?h k = (fds_nth f' k / k powr a) *\<^sub>R (\\<^sub>an. (-ln (real k) * (s - a)) ^ n / fact n)" by (subst infsetsum_scaleR_right [symmetric], rule infsetsum_cong) (simp_all add: scaleR_conv_of_real f_f' power_minus' power_mult_distrib divide_simps) also have "(\\<^sub>an. (-ln (real k) * (s - a)) ^ n / fact n) = exp (-ln (real k) * (s - a))" using exp_converges[of "-ln k * (s - a)"] exp_converges[of "norm (-ln k * (s - a))"] by (subst infsetsum_nat') (auto simp: abs_summable_on_nat_iff' sums_iff scaleR_conv_of_real divide_simps norm_divide norm_mult norm_power) also have "(fds_nth f' k / k powr a) *\<^sub>R \ = fds_nth f k / nat_power k s" by (auto simp: scaleR_conv_of_real f_f' powr_def exp_minus field_simps exp_of_real [symmetric] exp_diff) finally show ?case . qed finally have "fds_abs_converges f s" by (simp add: fds_abs_converges_def abs_summable_on_nat_iff') thus ?thesis by (intro exI[of _ "(c - \ / 2)"]) (auto simp: s_def a_def \) qed theorem pringsheim_landau: fixes c :: real and f :: "complex fds" assumes "nonneg_dirichlet_series f" assumes abscissa: "abs_conv_abscissa f = c" assumes g: "\s. s \ A \ Re s > c \ g s = eval_fds f s" assumes "g holomorphic_on A" "open A" "c \ A" shows False proof - have "\x complex_of_real x \ 1" unfolding abs_conv_abscissa_def by (intro Inf_lower) (auto simp: image_iff intro!: exI[of _ "of_real x"]) also have "\ < abs_conv_abscissa f" using assms x by simp finally show False by simp qed corollary entire_continuation_imp_abs_conv_abscissa_MInfty: assumes "nonneg_dirichlet_series f" assumes c: "c \ abs_conv_abscissa f" assumes g: "\s. Re s > c \ g s = eval_fds f s" assumes holo: "g holomorphic_on UNIV" shows "abs_conv_abscissa f = -\" proof (rule ccontr) assume "abs_conv_abscissa f \ -\" with c obtain a where abscissa [simp]: "abs_conv_abscissa f = ereal a" by (cases "abs_conv_abscissa f") auto show False proof (rule pringsheim_landau[OF assms(1) abscissa _ holo]) fix s assume s: "Re s > a" show "g s = eval_fds f s" proof (rule sym, rule analytic_continuation_open[of _ _ _ g]) show "g holomorphic_on {s. Re s > a}" by (rule holomorphic_on_subset[OF holo]) auto from assms show "{s. Re s > c} \ {s. Re s > a}" by auto next have "conv_abscissa f \ abs_conv_abscissa f" by (rule conv_le_abs_conv_abscissa) also have "\ = ereal a" by simp finally show "eval_fds f holomorphic_on {s. Re s > a}" by (intro holomorphic_intros) (auto intro: le_less_trans) qed (insert assms s, auto intro!: exI[of _ "of_real (c + 1)"] open_halfspace_Re_gt convex_connected convex_halfspace_Re_gt) qed auto qed subsection \Convergence of the $\zeta$ and M\"{o}bius $\mu$ series\ lemma fds_abs_summable_zeta_real_iff [simp]: "fds_abs_converges fds_zeta s \ s > (1 :: real)" proof - have "fds_abs_converges fds_zeta s \ summable (\n. real n powr -s)" unfolding fds_abs_converges_def by (intro summable_cong always_eventually) (auto simp: fds_nth_zeta powr_minus divide_simps) also have "\ \ s > 1" by (simp add: summable_real_powr_iff) finally show ?thesis . qed lemma fds_abs_summable_zeta_real: "s > (1 :: real) \ fds_abs_converges fds_zeta s" by simp lemma fds_abs_converges_moebius_mu_real: assumes "s > (1 :: real)" shows "fds_abs_converges (fds moebius_mu) s" unfolding fds_abs_converges_def proof (rule summable_comparison_test, intro exI allI impI) fix n :: nat show "norm (norm (fds_nth (fds moebius_mu) n / nat_power n s)) \ n powr (-s)" by (simp add: powr_minus divide_simps abs_moebius_mu_le) next from assms show "summable (\n. real n powr -s)" by (simp add: summable_real_powr_iff) qed subsection \Application to the M\"{o}bius $\mu$ function\ lemma inverse_squares_sums': "(\n. 1 / real n ^ 2) sums (pi ^ 2 / 6)" using inverse_squares_sums sums_Suc_iff[of "\n. 1 / real n ^ 2" "pi^2 / 6"] by simp lemma norm_summable_moebius_over_square: "summable (\n. norm (moebius_mu n / real n ^ 2))" proof (subst summable_Suc_iff [symmetric], rule summable_comparison_test) show "summable (\n. 1 / real (Suc n) ^ 2)" using inverse_squares_sums by (simp add: sums_iff) qed (auto simp del: of_nat_Suc simp: field_simps abs_moebius_mu_le) lemma summable_moebius_over_square: "summable (\n. moebius_mu n / real n ^ 2)" using norm_summable_moebius_over_square by (rule summable_norm_cancel) lemma moebius_over_square_sums: "(\n. moebius_mu n / n\<^sup>2) sums (6 / pi\<^sup>2)" proof - have "1 = eval_fds (1 :: real fds) 2" by simp also have "(1 :: real fds) = fds_zeta * fds moebius_mu" by (rule fds_zeta_times_moebius_mu [symmetric]) also have "eval_fds \ 2 = eval_fds fds_zeta 2 * eval_fds (fds moebius_mu) 2" by (intro eval_fds_mult fds_abs_converges_moebius_mu_real) simp_all also have "\ = pi ^ 2 / 6 * (\n. moebius_mu n / (real n)\<^sup>2)" using inverse_squares_sums' by (simp add: eval_fds_at_numeral suminf_fds_zeta_aux sums_iff) finally have "(\n. moebius_mu n / (real n)\<^sup>2) = 6 / pi ^ 2" by (simp add: field_simps) with summable_moebius_over_square show ?thesis by (simp add: sums_iff) qed end diff --git a/thys/Dirichlet_Series/ROOT b/thys/Dirichlet_Series/ROOT --- a/thys/Dirichlet_Series/ROOT +++ b/thys/Dirichlet_Series/ROOT @@ -1,18 +1,18 @@ chapter AFP -session Dirichlet_Series (AFP) = "HOL-Analysis" + +session Dirichlet_Series (AFP) = "HOL-Complex_Analysis" + options [timeout = 1200] sessions "HOL-Analysis" "HOL-Number_Theory" "HOL-Computational_Algebra" "HOL-Real_Asymp" "Polynomial_Factorization" "Landau_Symbols" "Euler_MacLaurin" theories Arithmetic_Summatory_Asymptotics Dirichlet_Efficient_Code document_files "root.tex" "root.bib" diff --git a/thys/E_Transcendental/E_Transcendental.thy b/thys/E_Transcendental/E_Transcendental.thy --- a/thys/E_Transcendental/E_Transcendental.thy +++ b/thys/E_Transcendental/E_Transcendental.thy @@ -1,687 +1,687 @@ (* File: E_Transcendental.thy Author: Manuel Eberl A proof that e (Euler's number) is transcendental. Could possibly be extended to a transcendence proof for pi or the very general Lindemann-Weierstrass theorem. *) section \Proof of the Transcendence of $e$\ theory E_Transcendental imports - "HOL-Analysis.Analysis" + "HOL-Complex_Analysis.Complex_Analysis" "HOL-Number_Theory.Number_Theory" "HOL-Computational_Algebra.Polynomial" begin (* TODO: Lots of stuff to move to the distribution *) subsection \Various auxiliary facts\ lemma fact_dvd_pochhammer: assumes "m \ n + 1" shows "fact m dvd pochhammer (int n - int m + 1) m" proof - have "(real n gchoose m) * fact m = of_int (pochhammer (int n - int m + 1) m)" by (simp add: gbinomial_pochhammer' pochhammer_of_int [symmetric]) also have "(real n gchoose m) * fact m = of_int (int (n choose m) * fact m)" by (simp add: binomial_gbinomial) finally have "int (n choose m) * fact m = pochhammer (int n - int m + 1) m" by (subst (asm) of_int_eq_iff) from this [symmetric] show ?thesis by simp qed lemma of_nat_eq_1_iff [simp]: "of_nat x = (1 :: 'a :: semiring_char_0) \ x = 1" by (fact of_nat_eq_1_iff) lemma prime_elem_int_not_dvd_neg1_power: "prime_elem (p :: int) \ \p dvd (-1) ^ n" by (rule notI, frule (1) prime_elem_dvd_power, cases "p \ 0") (auto simp: prime_elem_def) lemma nat_fact [simp]: "nat (fact n) = fact n" by (subst of_nat_fact [symmetric]) (rule nat_int) lemma prime_dvd_fact_iff_int: "p dvd fact n \ p \ int n" if "prime p" using that prime_dvd_fact_iff [of "nat \p\" n] by auto (simp add: prime_ge_0_int) lemma filterlim_minus_nat_at_top: "filterlim (\n. n - k :: nat) at_top at_top" proof - have "sequentially = filtermap (\n. n + k) at_top" by (auto simp: filter_eq_iff eventually_filtermap) also have "filterlim (\n. n - k :: nat) at_top \" by (simp add: filterlim_filtermap filterlim_ident) finally show ?thesis . qed lemma power_over_fact_tendsto_0: "(\n. (x :: real) ^ n / fact n) \ 0" using summable_exp[of x] by (intro summable_LIMSEQ_zero) (simp add: sums_iff field_simps) lemma power_over_fact_tendsto_0': "(\n. c * (x :: real) ^ n / fact n) \ 0" using tendsto_mult[OF tendsto_const[of c] power_over_fact_tendsto_0[of x]] by simp subsection \Lifting integer polynomials\ lift_definition of_int_poly :: "int poly \ 'a :: comm_ring_1 poly" is "\g x. of_int (g x)" by (auto elim: eventually_mono) lemma coeff_of_int_poly [simp]: "coeff (of_int_poly p) n = of_int (coeff p n)" by transfer' simp lemma of_int_poly_0 [simp]: "of_int_poly 0 = 0" by transfer (simp add: fun_eq_iff) lemma of_int_poly_pCons [simp]: "of_int_poly (pCons c p) = pCons (of_int c) (of_int_poly p)" by transfer' (simp add: fun_eq_iff split: nat.splits) lemma of_int_poly_smult [simp]: "of_int_poly (smult c p) = smult (of_int c) (of_int_poly p)" by transfer simp lemma of_int_poly_1 [simp]: "of_int_poly 1 = 1" by (simp add: one_pCons) lemma of_int_poly_add [simp]: "of_int_poly (p + q) = of_int_poly p + of_int_poly q" by transfer' (simp add: fun_eq_iff) lemma of_int_poly_mult [simp]: "of_int_poly (p * q) = (of_int_poly p * of_int_poly q)" by (induction p) simp_all lemma of_int_poly_sum [simp]: "of_int_poly (sum f A) = sum (\x. of_int_poly (f x)) A" by (induction A rule: infinite_finite_induct) simp_all lemma of_int_poly_prod [simp]: "of_int_poly (prod f A) = prod (\x. of_int_poly (f x)) A" by (induction A rule: infinite_finite_induct) simp_all lemma of_int_poly_power [simp]: "of_int_poly (p ^ n) = of_int_poly p ^ n" by (induction n) simp_all lemma of_int_poly_monom [simp]: "of_int_poly (monom c n) = monom (of_int c) n" by transfer (simp add: fun_eq_iff) lemma poly_of_int_poly [simp]: "poly (of_int_poly p) (of_int x) = of_int (poly p x)" by (induction p) simp_all lemma poly_of_int_poly_of_nat [simp]: "poly (of_int_poly p) (of_nat x) = of_int (poly p (int x))" by (induction p) simp_all lemma poly_of_int_poly_0 [simp]: "poly (of_int_poly p) 0 = of_int (poly p 0)" by (induction p) simp_all lemma poly_of_int_poly_1 [simp]: "poly (of_int_poly p) 1 = of_int (poly p 1)" by (induction p) simp_all lemma poly_of_int_poly_of_real [simp]: "poly (of_int_poly p) (of_real x) = of_real (poly (of_int_poly p) x)" by (induction p) simp_all lemma of_int_poly_eq_iff [simp]: "of_int_poly p = (of_int_poly q :: 'a :: {comm_ring_1, ring_char_0} poly) \ p = q" by (simp add: poly_eq_iff) lemma of_int_poly_eq_0_iff [simp]: "of_int_poly p = (0 :: 'a :: {comm_ring_1, ring_char_0} poly) \ p = 0" using of_int_poly_eq_iff[of p 0] by (simp del: of_int_poly_eq_iff) lemma degree_of_int_poly [simp]: "degree (of_int_poly p :: 'a :: {comm_ring_1, ring_char_0} poly) = degree p" by (simp add: degree_def) lemma pderiv_of_int_poly [simp]: "pderiv (of_int_poly p) = of_int_poly (pderiv p)" by (induction p) (simp_all add: pderiv_pCons) lemma higher_pderiv_of_int_poly [simp]: "(pderiv ^^ n) (of_int_poly p) = of_int_poly ((pderiv ^^ n) p)" by (induction n) simp_all lemma int_polyE: assumes "\n. coeff (p :: 'a :: {comm_ring_1, ring_char_0} poly) n \ \" obtains p' where "p = of_int_poly p'" proof - from assms have "\n. \c. coeff p n = of_int c" by (auto simp: Ints_def) hence "\c. \n. of_int (c n) = coeff p n" by (simp add: choice_iff eq_commute) then obtain c where c: "of_int (c n) = coeff p n" for n by blast have [simp]: "coeff (Abs_poly c) = c" proof (rule poly.Abs_poly_inverse, clarify) have "eventually (\n. n > degree p) at_top" by (rule eventually_gt_at_top) hence "eventually (\n. coeff p n = 0) at_top" by eventually_elim (simp add: coeff_eq_0) thus "eventually (\n. c n = 0) cofinite" by (simp add: c [symmetric] cofinite_eq_sequentially) qed have "p = of_int_poly (Abs_poly c)" by (rule poly_eqI) (simp add: c) thus ?thesis by (rule that) qed subsection \General facts about polynomials\ lemma pderiv_power: "pderiv (p ^ n) = smult (of_nat n) (p ^ (n - 1) * pderiv p)" by (cases n) (simp_all add: pderiv_power_Suc del: power_Suc) lemma degree_prod_sum_eq: "(\x. x \ A \ f x \ 0) \ degree (prod f A :: 'a :: idom poly) = (\x\A. degree (f x))" by (induction A rule: infinite_finite_induct) (auto simp: degree_mult_eq) lemma pderiv_monom: "pderiv (monom c n) = monom (of_nat n * c) (n - 1)" by (cases n) (simp_all add: monom_altdef pderiv_power_Suc pderiv_smult pderiv_pCons mult_ac del: power_Suc) lemma power_poly_const [simp]: "[:c:] ^ n = [:c ^ n:]" by (induction n) (simp_all add: power_commutes) lemma monom_power: "monom c n ^ k = monom (c ^ k) (n * k)" by (induction k) (simp_all add: mult_monom) lemma coeff_higher_pderiv: "coeff ((pderiv ^^ m) f) n = pochhammer (of_nat (Suc n)) m * coeff f (n + m)" by (induction m arbitrary: n) (simp_all add: coeff_pderiv pochhammer_rec algebra_simps) lemma higher_pderiv_add: "(pderiv ^^ n) (p + q) = (pderiv ^^ n) p + (pderiv ^^ n) q" by (induction n arbitrary: p q) (simp_all del: funpow.simps add: funpow_Suc_right pderiv_add) lemma higher_pderiv_smult: "(pderiv ^^ n) (smult c p) = smult c ((pderiv ^^ n) p)" by (induction n arbitrary: p) (simp_all del: funpow.simps add: funpow_Suc_right pderiv_smult) lemma higher_pderiv_0 [simp]: "(pderiv ^^ n) 0 = 0" by (induction n) simp_all lemma higher_pderiv_monom: "m \ n + 1 \ (pderiv ^^ m) (monom c n) = monom (pochhammer (int n - int m + 1) m * c) (n - m)" proof (induction m arbitrary: c n) case (Suc m) thus ?case by (cases n) (simp_all del: funpow.simps add: funpow_Suc_right pderiv_monom pochhammer_rec' Suc.IH) qed simp_all lemma higher_pderiv_monom_eq_zero: "m > n + 1 \ (pderiv ^^ m) (monom c n) = 0" proof (induction m arbitrary: c n) case (Suc m) thus ?case by (cases n) (simp_all del: funpow.simps add: funpow_Suc_right pderiv_monom pochhammer_rec' Suc.IH) qed simp_all lemma higher_pderiv_sum: "(pderiv ^^ n) (sum f A) = (\x\A. (pderiv ^^ n) (f x))" by (induction A rule: infinite_finite_induct) (simp_all add: higher_pderiv_add) lemma fact_dvd_higher_pderiv: "[:fact n :: int:] dvd (pderiv ^^ n) p" proof - have "[:fact n:] dvd (pderiv ^^ n) (monom c k)" for c :: int and k :: nat by (cases "n \ k + 1") (simp_all add: higher_pderiv_monom higher_pderiv_monom_eq_zero fact_dvd_pochhammer const_poly_dvd_iff) hence "[:fact n:] dvd (pderiv ^^ n) (\k\degree p. monom (coeff p k) k)" by (simp_all add: higher_pderiv_sum dvd_sum) thus ?thesis by (simp add: poly_as_sum_of_monoms) qed lemma fact_dvd_poly_higher_pderiv_aux: "(fact n :: int) dvd poly ((pderiv ^^ n) p) x" proof - have "[:fact n:] dvd (pderiv ^^ n) p" by (rule fact_dvd_higher_pderiv) then obtain q where "(pderiv ^^ n) p = [:fact n:] * q" by (erule dvdE) thus ?thesis by simp qed lemma fact_dvd_poly_higher_pderiv_aux': "m \ n \ (fact m :: int) dvd poly ((pderiv ^^ n) p) x" by (rule dvd_trans[OF fact_dvd fact_dvd_poly_higher_pderiv_aux]) simp_all lemma algebraicE': assumes "algebraic (x :: 'a :: field_char_0)" obtains p where "p \ 0" "poly (of_int_poly p) x = 0" proof - from assms obtain q where "\i. coeff q i \ \" "q \ 0" "poly q x = 0" by (erule algebraicE) moreover from this(1) obtain q' where "q = of_int_poly q'" by (erule int_polyE) ultimately show ?thesis by (intro that[of q']) simp_all qed lemma algebraicE'_nonzero: assumes "algebraic (x :: 'a :: field_char_0)" "x \ 0" obtains p where "p \ 0" "coeff p 0 \ 0" "poly (of_int_poly p) x = 0" proof - from assms(1) obtain p where p: "p \ 0" "poly (of_int_poly p) x = 0" by (erule algebraicE') define n :: nat where "n = order 0 p" have "monom 1 n dvd p" by (simp add: monom_1_dvd_iff p n_def) then obtain q where q: "p = monom 1 n * q" by (erule dvdE) from p have "q \ 0" "poly (of_int_poly q) x = 0" by (auto simp: q poly_monom assms(2)) moreover from this have "order 0 p = n + order 0 q" by (simp add: q order_mult) hence "order 0 q = 0" by (simp add: n_def) with \q \ 0\ have "poly q 0 \ 0" by (simp add: order_root) ultimately show ?thesis using that[of q] by (auto simp: poly_0_coeff_0) qed lemma algebraic_of_real_iff [simp]: "algebraic (of_real x :: 'a :: {real_algebra_1,field_char_0}) \ algebraic x" proof assume "algebraic (of_real x :: 'a)" then obtain p where "p \ 0" "poly (of_int_poly p) (of_real x :: 'a) = 0" by (erule algebraicE') hence "(of_int_poly p :: real poly) \ 0" "poly (of_int_poly p :: real poly) x = 0" by simp_all thus "algebraic x" by (intro algebraicI[of "of_int_poly p"]) simp_all next assume "algebraic x" then obtain p where "p \ 0" "poly (of_int_poly p) x = 0" by (erule algebraicE') hence "of_int_poly p \ (0 :: 'a poly)" "poly (of_int_poly p) (of_real x :: 'a) = 0" by simp_all thus "algebraic (of_real x)" by (intro algebraicI[of "of_int_poly p"]) simp_all qed subsection \Main proof\ lemma lindemann_weierstrass_integral: fixes u :: complex and f :: "complex poly" defines "df \ \n. (pderiv ^^ n) f" defines "m \ degree f" defines "I \ \f u. exp u * (\j\degree f. poly ((pderiv ^^ j) f) 0) - (\j\degree f. poly ((pderiv ^^ j) f) u)" shows "((\t. exp (u - t) * poly f t) has_contour_integral I f u) (linepath 0 u)" proof - note [derivative_intros] = exp_scaleR_has_vector_derivative_right vector_diff_chain_within let ?g = "\t. 1 - t" and ?f = "\t. -exp (t *\<^sub>R u)" have "((\t. exp ((1 - t) *\<^sub>R u) * u) has_integral (?f \ ?g) 1 - (?f \ ?g) 0) {0..1}" by (rule fundamental_theorem_of_calculus) (auto intro!: derivative_eq_intros simp del: o_apply) hence aux_integral: "((\t. exp (u - t *\<^sub>R u) * u) has_integral exp u - 1) {0..1}" by (simp add: algebra_simps) have "((\t. exp (u - t *\<^sub>R u) * u * poly f (t *\<^sub>R u)) has_integral I f u) {0..1}" unfolding df_def m_def proof (induction "degree f" arbitrary: f) case 0 then obtain c where c: "f = [:c:]" by (auto elim: degree_eq_zeroE) have "((\t. c * (exp (u - t *\<^sub>R u) * u)) has_integral c * (exp u - 1)) {0..1}" using aux_integral by (rule has_integral_mult_right) with c show ?case by (simp add: algebra_simps I_def) next case (Suc m) define df where "df = (\j. (pderiv ^^ j) f)" show ?case proof (rule integration_by_parts[OF bounded_bilinear_mult]) fix t :: real assume "t \ {0..1}" have "((?f \ ?g) has_vector_derivative exp (u - t *\<^sub>R u) * u) (at t)" by (auto intro!: derivative_eq_intros simp: algebra_simps simp del: o_apply) thus "((\t. -exp (u - t *\<^sub>R u)) has_vector_derivative exp (u - t *\<^sub>R u) * u) (at t)" by (simp add: algebra_simps o_def) next fix t :: real assume "t \ {0..1}" have "(poly f \ (\t. t *\<^sub>R u) has_vector_derivative u * poly (pderiv f) (t *\<^sub>R u)) (at t)" by (rule field_vector_diff_chain_at) (auto intro!: derivative_eq_intros) thus "((\t. poly f (t *\<^sub>R u)) has_vector_derivative u * poly (pderiv f) (t *\<^sub>R u)) (at t)" by (simp add: o_def) next from Suc(2) have m: "m = degree (pderiv f)" by (simp add: degree_pderiv) from Suc(1)[OF this] this have "((\t. exp (u - t *\<^sub>R u) * u * poly (pderiv f) (t *\<^sub>R u)) has_integral exp u * (\j=0..m. poly (df (Suc j)) 0) - (\j=0..m. poly (df (Suc j)) u)) {0..1}" by (simp add: df_def funpow_swap1 atMost_atLeast0 I_def) also have "(\j=0..m. poly (df (Suc j)) 0) = (\j=Suc 0..Suc m. poly (df j) 0)" by (rule sum.shift_bounds_cl_Suc_ivl [symmetric]) also have "\ = (\j=0..Suc m. poly (df j) 0) - poly f 0" by (subst (2) sum.atLeast_Suc_atMost) (simp_all add: df_def) also have "(\j=0..m. poly (df (Suc j)) u) = (\j=Suc 0..Suc m. poly (df j) u)" by (rule sum.shift_bounds_cl_Suc_ivl [symmetric]) also have "\ = (\j=0..Suc m. poly (df j) u) - poly f u" by (subst (2) sum.atLeast_Suc_atMost) (simp_all add: df_def) finally have "((\t. - (exp (u - t *\<^sub>R u) * u * poly (pderiv f) (t *\<^sub>R u))) has_integral -(exp u * ((\j = 0..Suc m. poly (df j) 0) - poly f 0) - ((\j = 0..Suc m. poly (df j) u) - poly f u))) {0..1}" (is "(_ has_integral ?I) _") by (rule has_integral_neg) also have "?I = - exp (u - 1 *\<^sub>R u) * poly f (1 *\<^sub>R u) - - exp (u - 0 *\<^sub>R u) * poly f (0 *\<^sub>R u) - I f u" by (simp add: df_def algebra_simps Suc(2) atMost_atLeast0 I_def) finally show "((\t. - exp (u - t *\<^sub>R u) * (u * poly (pderiv f) (t *\<^sub>R u))) has_integral \) {0..1}" by (simp add: algebra_simps) qed (auto intro!: continuous_intros) qed thus ?thesis by (simp add: has_contour_integral_linepath algebra_simps) qed locale lindemann_weierstrass_aux = fixes f :: "complex poly" begin definition I :: "complex \ complex" where "I u = exp u * (\j\degree f. poly ((pderiv ^^ j) f) 0) - (\j\degree f. poly ((pderiv ^^ j) f) u)" lemma lindemann_weierstrass_integral_bound: fixes u :: complex assumes "C \ 0" "\t. t \ closed_segment 0 u \ norm (poly f t) \ C" shows "norm (I u) \ norm u * exp (norm u) * C" proof - have "I u = contour_integral (linepath 0 u) (\t. exp (u - t) * poly f t)" using contour_integral_unique[OF lindemann_weierstrass_integral[of u f]] unfolding I_def .. also have "norm \ \ exp (norm u) * C * norm (u - 0)" proof (intro contour_integral_bound_linepath) fix t assume t: "t \ closed_segment 0 u" then obtain s where s: "s \ {0..1}" "t = s *\<^sub>R u" by (auto simp: closed_segment_def) hence "s * norm u \ 1 * norm u" by (intro mult_right_mono) simp_all with s have norm_t: "norm t \ norm u" by auto from s have "Re u - Re t = (1 - s) * Re u" by (simp add: algebra_simps) also have "\ \ norm u" proof (cases "Re u \ 0") case True with \s \ {0..1}\ have "(1 - s) * Re u \ 1 * Re u" by (intro mult_right_mono) simp_all also have "Re u \ norm u" by (rule complex_Re_le_cmod) finally show ?thesis by simp next case False with \s \ {0..1}\ have "(1 - s) * Re u \ 0" by (intro mult_nonneg_nonpos) simp_all also have "\ \ norm u" by simp finally show ?thesis . qed finally have "exp (Re u - Re t) \ exp (norm u)" by simp hence "exp (Re u - Re t) * norm (poly f t) \ exp (norm u) * C" using assms t norm_t by (intro mult_mono) simp_all thus "norm (exp (u - t) * poly f t) \ exp (norm u) * C" by (simp add: norm_mult exp_diff norm_divide field_simps) qed (auto simp: intro!: mult_nonneg_nonneg contour_integrable_continuous_linepath continuous_intros assms) finally show ?thesis by (simp add: mult_ac) qed end lemma poly_higher_pderiv_aux1: fixes c :: "'a :: idom" assumes "k < n" shows "poly ((pderiv ^^ k) ([:-c, 1:] ^ n * p)) c = 0" using assms proof (induction k arbitrary: n p) case (Suc k n p) from Suc.prems obtain n' where n: "n = Suc n'" by (cases n) auto from Suc.prems n have "k < n'" by simp have "(pderiv ^^ Suc k) ([:- c, 1:] ^ n * p) = (pderiv ^^ k) ([:- c, 1:] ^ n * pderiv p + [:- c, 1:] ^ n' * smult (of_nat n) p)" by (simp only: funpow_Suc_right o_def pderiv_mult n pderiv_power_Suc, simp only: n [symmetric]) (simp add: pderiv_pCons mult_ac) also from Suc.prems \k < n'\ have "poly \ c = 0" by (simp add: higher_pderiv_add Suc.IH del: mult_smult_right) finally show ?case . qed simp_all lemma poly_higher_pderiv_aux1': fixes c :: "'a :: idom" assumes "k < n" "[:-c, 1:] ^ n dvd p" shows "poly ((pderiv ^^ k) p) c = 0" proof - from assms(2) obtain q where "p = [:-c, 1:] ^ n * q" by (elim dvdE) also from assms(1) have "poly ((pderiv ^^ k) \) c = 0" by (rule poly_higher_pderiv_aux1) finally show ?thesis . qed lemma poly_higher_pderiv_aux2: fixes c :: "'a :: {idom, semiring_char_0}" shows "poly ((pderiv ^^ n) ([:-c, 1:] ^ n * p)) c = fact n * poly p c" proof (induction n arbitrary: p) case (Suc n p) have "(pderiv ^^ Suc n) ([:- c, 1:] ^ Suc n * p) = (pderiv ^^ n) ([:- c, 1:] ^ Suc n * pderiv p) + (pderiv ^^ n) ([:- c, 1:] ^ n * smult (1 + of_nat n) p)" by (simp del: funpow.simps power_Suc add: funpow_Suc_right pderiv_mult pderiv_power_Suc higher_pderiv_add pderiv_pCons mult_ac) also have "[:- c, 1:] ^ Suc n * pderiv p = [:- c, 1:] ^ n * ([:-c, 1:] * pderiv p)" by (simp add: algebra_simps) finally show ?case by (simp add: Suc.IH del: mult_smult_right power_Suc) qed simp_all lemma poly_higher_pderiv_aux3: fixes c :: "'a :: {idom,semiring_char_0}" assumes "k \ n" shows "\q. poly ((pderiv ^^ k) ([:-c, 1:] ^ n * p)) c = fact n * poly q c" using assms proof (induction k arbitrary: n p) case (Suc k n p) show ?case proof (cases n) fix n' assume n: "n = Suc n'" have "poly ((pderiv ^^ Suc k) ([:-c, 1:] ^ n * p)) c = poly ((pderiv ^^ k) ([:- c, 1:] ^ n * pderiv p)) c + of_nat n * poly ((pderiv ^^ k) ([:-c, 1:] ^ n' * p)) c" by (simp del: funpow.simps power_Suc add: funpow_Suc_right pderiv_power_Suc pderiv_mult n pderiv_pCons higher_pderiv_add mult_ac higher_pderiv_smult) also have "\q1. poly ((pderiv ^^ k) ([:-c, 1:] ^ n * pderiv p)) c = fact n * poly q1 c" using Suc.prems Suc.IH[of n "pderiv p"] by (cases "n' = k") (auto simp: n poly_higher_pderiv_aux1 simp del: power_Suc of_nat_Suc intro: exI[of _ "0::'a poly"]) then obtain q1 where "poly ((pderiv ^^ k) ([:-c, 1:] ^ n * pderiv p)) c = fact n * poly q1 c" .. also from Suc.IH[of n' p] Suc.prems obtain q2 where "poly ((pderiv ^^ k) ([:-c, 1:] ^ n' * p)) c = fact n' * poly q2 c" by (auto simp: n) finally show ?case by (auto intro!: exI[of _ "q1 + q2"] simp: n algebra_simps) qed auto qed auto lemma poly_higher_pderiv_aux3': fixes c :: "'a :: {idom, semiring_char_0}" assumes "k \ n" "[:-c, 1:] ^ n dvd p" shows "fact n dvd poly ((pderiv ^^ k) p) c" proof - from assms(2) obtain q where "p = [:-c, 1:] ^ n * q" by (elim dvdE) with poly_higher_pderiv_aux3[OF assms(1), of c q] show ?thesis by auto qed lemma e_transcendental_aux_bound: obtains C where "C \ 0" "\x. x \ closed_segment 0 (of_nat n) \ norm (\k\{1..n}. (x - of_nat k :: complex)) \ C" proof - let ?f = "\x. (\k\{1..n}. (x - of_nat k))" define C where "C = max 0 (Sup (cmod ` ?f ` closed_segment 0 (of_nat n)))" have "C \ 0" by (simp add: C_def) moreover { fix x :: complex assume "x \ closed_segment 0 (of_nat n)" hence "cmod (?f x) \ Sup ((cmod \ ?f) ` closed_segment 0 (of_nat n))" by (intro cSup_upper bounded_imp_bdd_above compact_imp_bounded compact_continuous_image) (auto intro!: continuous_intros) also have "\ \ C" by (simp add: C_def image_comp) finally have "cmod (?f x) \ C" . } ultimately show ?thesis by (rule that) qed theorem e_transcendental_complex: "\ algebraic (exp 1 :: complex)" proof assume "algebraic (exp 1 :: complex)" then obtain q :: "int poly" where q: "q \ 0" "coeff q 0 \ 0" "poly (of_int_poly q) (exp 1 :: complex) = 0" by (elim algebraicE'_nonzero) simp_all define n :: nat where "n = degree q" from q have [simp]: "n \ 0" by (intro notI) (auto simp: n_def elim!: degree_eq_zeroE) define qmax where "qmax = Max (insert 0 (abs ` set (coeffs q)))" have qmax_nonneg [simp]: "qmax \ 0" by (simp add: qmax_def) have qmax: "\coeff q k\ \ qmax" for k by (cases "k \ degree q") (auto simp: qmax_def coeff_eq_0 coeffs_def simp del: upt_Suc intro: Max.coboundedI) obtain C where C: "C \ 0" "\x. x \ closed_segment 0 (of_nat n) \ norm (\k\{1..n}. (x - of_nat k :: complex)) \ C" by (erule e_transcendental_aux_bound) define E where "E = (1 + real n) * real_of_int qmax * real n * exp (real n) / real n" define F where "F = real n * C" have ineq: "fact (p - 1) \ E * F ^ p" if p: "prime p" "p > n" "p > abs (coeff q 0)" for p proof - from p(1) have p_pos: "p > 0" by (simp add: prime_gt_0_nat) define f :: "int poly" where "f = monom 1 (p - 1) * (\k\{1..n}. [:-of_nat k, 1:] ^ p)" have poly_f: "poly (of_int_poly f) x = x ^ (p - 1) * (\k\{1..n}. (x - of_nat k)) ^ p" for x :: complex by (simp add: f_def poly_prod poly_monom prod_power_distrib) define m :: nat where "m = degree f" from p_pos have m: "m = (n + 1) * p - 1" by (simp add: m_def f_def degree_mult_eq degree_monom_eq degree_prod_sum_eq degree_linear_power) define M :: int where "M = (- 1) ^ (n * p) * fact n ^ p" with p have p_not_dvd_M: "\int p dvd M" by (auto simp: M_def prime_elem_int_not_dvd_neg1_power prime_dvd_power_iff prime_gt_0_nat prime_dvd_fact_iff_int prime_dvd_mult_iff) interpret lindemann_weierstrass_aux "of_int_poly f" . define J :: complex where "J = (\k\n. of_int (coeff q k) * I (of_nat k))" define idxs where "idxs = ({..n}\{..m}) - {(0, p - 1)}" hence "J = (\k\n. of_int (coeff q k) * exp 1 ^ k) * (\n\m. of_int (poly ((pderiv ^^ n) f) 0)) - of_int (\k\n. \n\m. coeff q k * poly ((pderiv ^^ n) f) (int k))" by (simp add: J_def I_def algebra_simps sum_subtractf sum_distrib_left m_def exp_of_nat_mult [symmetric]) also have "(\k\n. of_int (coeff q k) * exp 1 ^ k) = poly (of_int_poly q) (exp 1 :: complex)" by (simp add: poly_altdef n_def) also have "\ = 0" by fact finally have "J = of_int (-(\(k,n)\{..n}\{..m}. coeff q k * poly ((pderiv ^^ n) f) (int k)))" by (simp add: sum.cartesian_product) also have "{..n}\{..m} = insert (0, p - 1) idxs" by (auto simp: m idxs_def) also have "-(\(k,n)\\. coeff q k * poly ((pderiv ^^ n) f) (int k)) = - (coeff q 0 * poly ((pderiv ^^ (p - 1)) f) 0) - (\(k, n)\idxs. coeff q k * poly ((pderiv ^^ n) f) (of_nat k))" by (subst sum.insert) (simp_all add: idxs_def) also have "coeff q 0 * poly ((pderiv ^^ (p - 1)) f) 0 = coeff q 0 * M * fact (p - 1)" proof - have "f = [:-0, 1:] ^ (p - 1) * (\k = 1..n. [:- of_nat k, 1:] ^ p)" by (simp add: f_def monom_altdef) also have "poly ((pderiv ^^ (p - 1)) \) 0 = fact (p - 1) * poly (\k = 1..n. [:- of_nat k, 1:] ^ p) 0" by (rule poly_higher_pderiv_aux2) also have "poly (\k = 1..n. [:- of_nat k :: int, 1:] ^ p) 0 = (-1)^(n*p) * fact n ^ p" by (induction n) (simp_all add: prod.nat_ivl_Suc' power_mult_distrib mult_ac power_minus' power_add del: of_nat_Suc) finally show ?thesis by (simp add: mult_ac M_def) qed also have "\N. (\(k, n)\idxs. coeff q k * poly ((pderiv ^^ n) f) (int k)) = fact p * N" proof - have "\(k, n)\idxs. fact p dvd poly ((pderiv ^^ n) f) (of_nat k)" proof clarify fix k j assume idxs: "(k, j) \ idxs" then consider "k = 0" "j < p - 1" | "k = 0" "j > p - 1" | "k \ 0" "j < p" | "k \ 0" "j \ p" by (fastforce simp: idxs_def) thus "fact p dvd poly ((pderiv ^^ j) f) (of_nat k)" proof cases case 1 thus ?thesis by (simp add: f_def poly_higher_pderiv_aux1' monom_altdef) next case 2 thus ?thesis by (simp add: f_def poly_higher_pderiv_aux3' monom_altdef fact_dvd_poly_higher_pderiv_aux') next case 3 thus ?thesis unfolding f_def by (subst poly_higher_pderiv_aux1'[of _ p]) (insert idxs, auto simp: idxs_def intro!: dvd_mult) next case 4 thus ?thesis unfolding f_def by (intro poly_higher_pderiv_aux3') (insert idxs, auto intro!: dvd_mult simp: idxs_def) qed qed hence "fact p dvd (\(k, n)\idxs. coeff q k * poly ((pderiv ^^ n) f) (int k))" by (auto intro!: dvd_sum dvd_mult simp del: of_int_fact) thus ?thesis by (blast elim: dvdE) qed then guess N .. note N = this also from p have "- (coeff q 0 * M * fact (p - 1)) - fact p * N = - fact (p - 1) * (coeff q 0 * M + p * N)" by (subst fact_reduce[of p]) (simp_all add: algebra_simps) finally have J: "J = -of_int (fact (p - 1) * (coeff q 0 * M + p * N))" by simp from p q(2) have "\p dvd coeff q 0 * M + p * N" by (auto simp: dvd_add_left_iff p_not_dvd_M prime_dvd_fact_iff_int prime_dvd_mult_iff dest: dvd_imp_le_int) hence "coeff q 0 * M + p * N \ 0" by (intro notI) simp_all hence "abs (coeff q 0 * M + p * N) \ 1" by simp hence "norm (of_int (coeff q 0 * M + p * N) :: complex) \ 1" by (simp only: norm_of_int) hence "fact (p - 1) * \ \ fact (p - 1) * 1" by (intro mult_left_mono) simp_all hence J_lower: "norm J \ fact (p - 1)" unfolding J norm_minus_cancel of_int_mult of_int_fact by (simp add: norm_mult) have "norm J \ (\k\n. norm (of_int (coeff q k) * I (of_nat k)))" unfolding J_def by (rule norm_sum) also have "\ \ (\k\n. of_int qmax * (real n * exp (real n) * real n ^ (p - 1) * C ^ p))" proof (intro sum_mono) fix k assume k: "k \ {..n}" have "n > 0" by (rule ccontr) simp { fix x :: complex assume x: "x \ closed_segment 0 (of_nat k)" then obtain t where t: "t \ 0" "t \ 1" "x = of_real t * of_nat k" by (auto simp: closed_segment_def scaleR_conv_of_real) hence "norm x = t * real k" by (simp add: norm_mult) also from \t \ 1\ k have *: "\ \ 1 * real n" by (intro mult_mono) simp_all finally have x': "norm x \ real n" by simp from t \n > 0\ * have x'': "x \ closed_segment 0 (of_nat n)" by (auto simp: closed_segment_def scaleR_conv_of_real field_simps intro!: exI[of _ "t * real k / real n"] ) have "norm (poly (of_int_poly f) x) = norm x ^ (p - 1) * cmod (\i = 1..n. x - i) ^ p" by (simp add: poly_f norm_mult norm_power) also from x x' x'' have "\ \ of_nat n ^ (p - 1) * C ^ p" by (intro mult_mono C power_mono) simp_all finally have "norm (poly (of_int_poly f) x) \ real n ^ (p - 1) * C ^ p" . } note A = this have "norm (I (of_nat k)) \ cmod (of_nat k) * exp (cmod (of_nat k)) * (of_nat n ^ (p - 1) * C ^ p)" by (intro lindemann_weierstrass_integral_bound[OF _ A] C mult_nonneg_nonneg zero_le_power) auto also have "\ \ cmod (of_nat n) * exp (cmod (of_nat n)) * (of_nat n ^ (p - 1) * C ^ p)" using k by (intro mult_mono zero_le_power mult_nonneg_nonneg C) simp_all finally show "cmod (of_int (coeff q k) * I (of_nat k)) \ of_int qmax * (real n * exp (real n) * real n ^ (p - 1) * C ^ p)" unfolding norm_mult by (intro mult_mono) (simp_all add: qmax of_int_abs [symmetric] del: of_int_abs) qed also have "\ = E * F ^ p" using p_pos by (simp add: power_diff power_mult_distrib E_def F_def) finally show "fact (p - 1) \ E * F ^ p" using J_lower by linarith qed have "(\n. E * F * F ^ (n - 1) / fact (n - 1)) \ 0" (is ?P) by (intro filterlim_compose[OF power_over_fact_tendsto_0' filterlim_minus_nat_at_top]) also have "?P \ (\n. E * F ^ n / fact (n - 1)) \ 0" by (intro filterlim_cong refl eventually_mono[OF eventually_gt_at_top[of "0::nat"]]) (auto simp: power_Suc [symmetric] simp del: power_Suc) finally have "eventually (\n. E * F ^ n / fact (n - 1) < 1) at_top" by (rule order_tendstoD) simp_all hence "eventually (\n. E * F ^ n < fact (n - 1)) at_top" by eventually_elim simp then obtain P where P: "\n. n \ P \ E * F ^ n < fact (n - 1)" by (auto simp: eventually_at_top_linorder) have "\p. prime p \ p > Max {nat (abs (coeff q 0)), n, P}" by (rule bigger_prime) then obtain p where "prime p" "p > Max {nat (abs (coeff q 0)), n, P}" by blast hence "int p > abs (coeff q 0)" "p > n" "p \ P" by auto with ineq[of p] \prime p\ have "fact (p - 1) \ E * F ^ p" by simp moreover from \p \ P\ have "fact (p - 1) > E * F ^ p" by (rule P) ultimately show False by linarith qed corollary e_transcendental_real: "\ algebraic (exp 1 :: real)" proof - have "\algebraic (exp 1 :: complex)" by (rule e_transcendental_complex) also have "(exp 1 :: complex) = of_real (exp 1)" using exp_of_real[of 1] by simp also have "algebraic \ \ algebraic (exp 1 :: real)" by simp finally show ?thesis . qed end diff --git a/thys/E_Transcendental/ROOT b/thys/E_Transcendental/ROOT --- a/thys/E_Transcendental/ROOT +++ b/thys/E_Transcendental/ROOT @@ -1,11 +1,11 @@ chapter AFP -session E_Transcendental (AFP) = "HOL-Analysis" + +session E_Transcendental (AFP) = "HOL-Complex_Analysis" + options [timeout = 600] sessions "HOL-Number_Theory" theories E_Transcendental document_files "root.tex" "root.bib" diff --git a/thys/Error_Function/Error_Function.thy b/thys/Error_Function/Error_Function.thy --- a/thys/Error_Function/Error_Function.thy +++ b/thys/Error_Function/Error_Function.thy @@ -1,655 +1,655 @@ (* File: Error_Function.thy Author: Manuel Eberl, TU München The complex and real error function (most results are on the real error function though) *) section \The complex and real error function\ theory Error_Function - imports "HOL-Analysis.Analysis" "HOL-Library.Landau_Symbols" + imports "HOL-Complex_Analysis.Complex_Analysis" "HOL-Library.Landau_Symbols" begin subsection \Auxiliary Facts\ lemma tendsto_sandwich_mono: assumes "(\n. f (real n)) \ (c::real)" assumes "eventually (\x. \y z. x \ y \ y \ z \ f y \ f z) at_top" shows "(f \ c) at_top" proof (rule tendsto_sandwich) from assms(2) obtain C where C: "\x y. C \ x \ x \ y \ f x \ f y" by (auto simp: eventually_at_top_linorder) show "eventually (\x. f (real (nat \x\)) \ f x) at_top" using eventually_gt_at_top[of "0::real"] eventually_gt_at_top[of "C+1::real"] by eventually_elim (rule C, linarith+) show "eventually (\x. f (real (Suc (nat \x\))) \ f x) at_top" using eventually_gt_at_top[of "0::real"] eventually_gt_at_top[of "C+1::real"] by eventually_elim (rule C, linarith+) qed (auto intro!: filterlim_compose[OF assms(1)] filterlim_compose[OF filterlim_nat_sequentially] filterlim_compose[OF filterlim_Suc] filterlim_floor_sequentially simp del: of_nat_Suc) lemma tendsto_sandwich_antimono: assumes "(\n. f (real n)) \ (c::real)" assumes "eventually (\x. \y z. x \ y \ y \ z \ f y \ f z) at_top" shows "(f \ c) at_top" proof (rule tendsto_sandwich) from assms(2) obtain C where C: "\x y. C \ x \ x \ y \ f x \ f y" by (auto simp: eventually_at_top_linorder) show "eventually (\x. f (real (nat \x\)) \ f x) at_top" using eventually_gt_at_top[of "0::real"] eventually_gt_at_top[of "C+1::real"] by eventually_elim (rule C, linarith+) show "eventually (\x. f (real (Suc (nat \x\))) \ f x) at_top" using eventually_gt_at_top[of "0::real"] eventually_gt_at_top[of "C+1::real"] by eventually_elim (rule C, linarith+) qed (auto intro!: filterlim_compose[OF assms(1)] filterlim_compose[OF filterlim_nat_sequentially] filterlim_compose[OF filterlim_Suc] filterlim_floor_sequentially simp del: of_nat_Suc) lemma has_bochner_integral_completion [intro]: fixes f :: "'a \ 'b::{banach, second_countable_topology}" shows "has_bochner_integral M f I \ has_bochner_integral (completion M) f I" by (auto simp: has_bochner_integral_iff integrable_completion integral_completion borel_measurable_integrable) lemma has_bochner_integral_imp_has_integral: "has_bochner_integral lebesgue (\x. indicator S x *\<^sub>R f x) I \ (f has_integral (I :: 'b :: euclidean_space)) S" using has_integral_set_lebesgue[of S f] by (simp add: has_bochner_integral_iff set_integrable_def set_lebesgue_integral_def) lemma has_bochner_integral_imp_has_integral': "has_bochner_integral lborel (\x. indicator S x *\<^sub>R f x) I \ (f has_integral (I :: 'b :: euclidean_space)) S" by (intro has_bochner_integral_imp_has_integral has_bochner_integral_completion) lemma has_bochner_integral_erf_aux: "has_bochner_integral lborel (\x. indicator {0..} x *\<^sub>R exp (- x\<^sup>2)) (sqrt pi / 2)" proof - let ?pI = "\f. (\\<^sup>+s. f (s::real) * indicator {0..} s \lborel)" let ?gauss = "\x. exp (- x\<^sup>2)" let ?I = "indicator {0<..} :: real \ real" let ?ff= "\x s. x * exp (- x\<^sup>2 * (1 + s\<^sup>2)) :: real" have *: "?pI ?gauss = (\\<^sup>+x. ?gauss x * ?I x \lborel)" by (intro nn_integral_cong_AE AE_I[where N="{0}"]) (auto split: split_indicator) have "?pI ?gauss * ?pI ?gauss = (\\<^sup>+x. \\<^sup>+s. ?gauss x * ?gauss s * ?I s * ?I x \lborel \lborel)" by (auto simp: nn_integral_cmult[symmetric] nn_integral_multc[symmetric] * ennreal_mult[symmetric] intro!: nn_integral_cong split: split_indicator) also have "\ = (\\<^sup>+x. \\<^sup>+s. ?ff x s * ?I s * ?I x \lborel \lborel)" proof (rule nn_integral_cong, cases) fix x :: real assume "x \ 0" then show "(\\<^sup>+s. ?gauss x * ?gauss s * ?I s * ?I x \lborel) = (\\<^sup>+s. ?ff x s * ?I s * ?I x \lborel)" by (subst nn_integral_real_affine[where t="0" and c="x"]) (auto simp: mult_exp_exp nn_integral_cmult[symmetric] field_simps zero_less_mult_iff ennreal_mult[symmetric] intro!: nn_integral_cong split: split_indicator) qed simp also have "... = \\<^sup>+s. \\<^sup>+x. ?ff x s * ?I s * ?I x \lborel \lborel" by (rule lborel_pair.Fubini'[symmetric]) auto also have "... = ?pI (\s. ?pI (\x. ?ff x s))" by (rule nn_integral_cong_AE) (auto intro!: nn_integral_cong_AE AE_I[where N="{0}"] split: split_indicator_asm) also have "\ = ?pI (\s. ennreal (1 / (2 * (1 + s\<^sup>2))))" proof (intro nn_integral_cong ennreal_mult_right_cong) fix s :: real show "?pI (\x. ?ff x s) = ennreal (1 / (2 * (1 + s\<^sup>2)))" proof (subst nn_integral_FTC_atLeast) have "((\a. - (exp (- ((1 + s\<^sup>2) * a\<^sup>2)) / (2 + 2 * s\<^sup>2))) \ (- (0 / (2 + 2 * s\<^sup>2)))) at_top" by (intro tendsto_intros filterlim_compose[OF exp_at_bot] filterlim_compose[OF filterlim_uminus_at_bot_at_top] filterlim_tendsto_pos_mult_at_top) (auto intro!: filterlim_at_top_mult_at_top[OF filterlim_ident filterlim_ident] simp: add_pos_nonneg power2_eq_square add_nonneg_eq_0_iff) then show "((\a. - (exp (- a\<^sup>2 - s\<^sup>2 * a\<^sup>2) / (2 + 2 * s\<^sup>2))) \ 0) at_top" by (simp add: field_simps) qed (auto intro!: derivative_eq_intros simp: field_simps add_nonneg_eq_0_iff) qed also have "... = ennreal (pi / 4)" proof (subst nn_integral_FTC_atLeast) show "((\a. arctan a / 2) \ (pi / 2) / 2 ) at_top" by (intro tendsto_intros) (simp_all add: tendsto_arctan_at_top) qed (auto intro!: derivative_eq_intros simp: add_nonneg_eq_0_iff field_simps power2_eq_square) finally have "?pI ?gauss^2 = pi / 4" by (simp add: power2_eq_square) then have "?pI ?gauss = sqrt (pi / 4)" using power_eq_iff_eq_base[of 2 "enn2real (?pI ?gauss)" "sqrt (pi / 4)"] by (cases "?pI ?gauss") (auto simp: power2_eq_square ennreal_mult[symmetric] ennreal_top_mult) also have "?pI ?gauss = (\\<^sup>+x. indicator {0..} x *\<^sub>R exp (- x\<^sup>2) \lborel)" by (intro nn_integral_cong) (simp split: split_indicator) also have "sqrt (pi / 4) = sqrt pi / 2" by (simp add: real_sqrt_divide) finally show ?thesis by (rule has_bochner_integral_nn_integral[rotated 3]) auto qed lemma has_integral_erf_aux: "((\t::real. exp (-t\<^sup>2)) has_integral (sqrt pi / 2)) {0..}" by (intro has_bochner_integral_imp_has_integral' has_bochner_integral_erf_aux) lemma contour_integrable_on_linepath_neg_exp_squared [simp, intro]: "(\t. exp (-(t^2))) contour_integrable_on linepath 0 z" by (auto intro!: contour_integrable_continuous_linepath continuous_intros) lemma holomorphic_on_chain: "g holomorphic_on t \ f holomorphic_on s \ f ` s \ t \ (\x. g (f x)) holomorphic_on s" using holomorphic_on_compose_gen[of f s g t] by (simp add: o_def) lemma holomorphic_on_chain_UNIV: "g holomorphic_on UNIV \ f holomorphic_on s\ (\x. g (f x)) holomorphic_on s" using holomorphic_on_compose_gen[of f s g UNIV] by (simp add: o_def) lemmas holomorphic_on_exp' [holomorphic_intros] = holomorphic_on_exp [THEN holomorphic_on_chain_UNIV] lemma leibniz_rule_field_derivative_real: fixes f::"'a::{real_normed_field, banach} \ real \ 'a" assumes fx: "\x t. x \ U \ t \ {a..b} \ ((\x. f x t) has_field_derivative fx x t) (at x within U)" assumes integrable_f2: "\x. x \ U \ (f x) integrable_on {a..b}" assumes cont_fx: "continuous_on (U \ {a..b}) (\(x, t). fx x t)" assumes U: "x0 \ U" "convex U" shows "((\x. integral {a..b} (f x)) has_field_derivative integral {a..b} (fx x0)) (at x0 within U)" using leibniz_rule_field_derivative[of U a b f fx x0] assms by simp lemma has_vector_derivative_linepath_within [derivative_intros]: assumes [derivative_intros]: "(f has_vector_derivative f') (at x within S)" "(g has_vector_derivative g') (at x within S)" "(h has_real_derivative h') (at x within S)" shows "((\x. linepath (f x) (g x) (h x)) has_vector_derivative (1 - h x) *\<^sub>R f' + h x *\<^sub>R g' - h' *\<^sub>R (f x - g x)) (at x within S)" unfolding linepath_def [abs_def] by (auto intro!: derivative_eq_intros simp: field_simps scaleR_diff_right) lemma has_field_derivative_linepath_within [derivative_intros]: assumes [derivative_intros]: "(f has_field_derivative f') (at x within S)" "(g has_field_derivative g') (at x within S)" "(h has_real_derivative h') (at x within S)" shows "((\x. linepath (f x) (g x) (h x)) has_field_derivative (1 - h x) *\<^sub>R f' + h x *\<^sub>R g' - h' *\<^sub>R (f x - g x)) (at x within S)" unfolding linepath_def [abs_def] by (auto intro!: derivative_eq_intros simp: field_simps scaleR_diff_right) lemma continuous_on_linepath' [continuous_intros]: assumes [continuous_intros]: "continuous_on A f" "continuous_on A g" "continuous_on A h" shows "continuous_on A (\x. linepath (f x) (g x) (h x))" using assms unfolding linepath_def by (auto intro!: continuous_intros) lemma contour_integral_has_field_derivative: assumes A: "open A" "convex A" "a \ A" "z \ A" assumes integrable: "\z. z \ A \ f contour_integrable_on linepath a z" assumes holo: "f holomorphic_on A" shows "((\z. contour_integral (linepath a z) f) has_field_derivative f z) (at z within B)" proof - have "(f has_field_derivative deriv f z) (at z)" if "z \ A" for z using that assms by (auto intro!: holomorphic_derivI) note [derivative_intros] = DERIV_chain2[OF this] note [continuous_intros] = continuous_on_compose2[OF holomorphic_on_imp_continuous_on [OF holo]] continuous_on_compose2[OF holomorphic_on_imp_continuous_on [OF holomorphic_deriv[OF holo]]] have [derivative_intros]: "((\x. linepath a x t) has_field_derivative of_real t) (at x within A)" for t x by (auto simp: linepath_def scaleR_conv_of_real intro!: derivative_eq_intros) have *: "linepath a b t \ A" if "a \ A" "b \ A" "t \ {0..1}" for a b t using that linepath_in_convex_hull[of a A b t] \convex A\ by (simp add: hull_same) have "((\z. integral {0..1} (\x. f (linepath a z x)) * (z - a)) has_field_derivative integral {0..1} (\t. deriv f (linepath a z t) * of_real t) * (z - a) + integral {0..1} (\x. f (linepath a z x))) (at z within A)" (is "(_ has_field_derivative ?I) _") by (rule derivative_eq_intros leibniz_rule_field_derivative_real)+ (insert assms, auto intro!: derivative_eq_intros leibniz_rule_field_derivative_real integrable_continuous_real continuous_intros simp: split_beta scaleR_conv_of_real *) also have "(\z. integral {0..1} (\x. f (linepath a z x)) * (z - a)) = (\z. contour_integral (linepath a z) f)" by (simp add: contour_integral_integral) also have "?I = integral {0..1} (\x. deriv f (linepath a z x) * of_real x * (z - a) + f (linepath a z x))" (is "_ = integral _ ?g") by (subst integral_mult_left [symmetric], subst integral_add [symmetric]) (insert assms, auto intro!: integrable_continuous_real continuous_intros simp: *) also have "(?g has_integral of_real 1 * f (linepath a z 1) - of_real 0 * f (linepath a z 0)) {0..1}" using * A by (intro fundamental_theorem_of_calculus) (auto intro!: derivative_eq_intros has_vector_derivative_real_field simp: linepath_def scaleR_conv_of_real) hence "integral {0..1} ?g = f (linepath a z 1)" by (simp add: has_integral_iff) also have "linepath a z 1 = z" by (simp add: linepath_def) also from \z \ A\ and \open A\ have "at z within A = at z" by (rule at_within_open) finally show ?thesis by (rule DERIV_subset) simp_all qed subsection \Definition of the error function\ definition erf_coeffs :: "nat \ real" where "erf_coeffs n = (if odd n then 2 / sqrt pi * (-1) ^ (n div 2) / (of_nat n * fact (n div 2)) else 0)" lemma summable_erf: fixes z :: "'a :: {real_normed_div_algebra, banach}" shows "summable (\n. of_real (erf_coeffs n) * z ^ n)" proof - define b where "b = (\n. 2 / sqrt pi * (if odd n then inverse (fact (n div 2)) else 0))" show ?thesis proof (rule summable_comparison_test[OF exI[of _ 1]], clarify) fix n :: nat assume n: "n \ 1" hence "norm (of_real (erf_coeffs n) * z ^ n) \ b n * norm z ^ n" unfolding norm_mult norm_power erf_coeffs_def b_def by (intro mult_right_mono) (auto simp: field_simps norm_divide abs_mult) thus "norm (of_real (erf_coeffs n) * z ^ n) \ b n * norm z ^ n" by (simp add: mult_ac) next have "summable (\n. (norm z * 2 / sqrt pi) * (inverse (fact n) * norm z ^ (2*n)))" (is "summable ?c") unfolding power_mult by (intro summable_mult summable_exp) also have "?c = (\n. b (2*n+1) * norm z ^ (2*n+1))" unfolding b_def by (auto simp: fun_eq_iff b_def) also have "summable \ \ summable (\n. b n * norm z ^ n)" using summable_mono_reindex [of "\n. 2*n+1"] by (intro summable_mono_reindex [of "\n. 2*n+1"]) (auto elim!: oddE simp: strict_mono_def b_def) finally show \ . qed qed definition erf :: "('a :: {real_normed_field, banach}) \ 'a" where "erf z = (\n. of_real (erf_coeffs n) * z ^ n)" lemma erf_converges: "(\n. of_real (erf_coeffs n) * z ^ n) sums erf z" using summable_erf by (simp add: sums_iff erf_def) lemma erf_0 [simp]: "erf 0 = 0" unfolding erf_def powser_zero by (simp add: erf_coeffs_def) lemma erf_minus [simp]: "erf (-z) = - erf z" unfolding erf_def by (subst suminf_minus [OF summable_erf, symmetric], rule suminf_cong) (simp_all add: erf_coeffs_def) lemma erf_of_real [simp]: "erf (of_real x) = of_real (erf x)" unfolding erf_def using summable_erf[of x] by (subst suminf_of_real) (simp_all add: summable_erf) lemma of_real_erf_numeral [simp]: "of_real (erf (numeral n)) = erf (numeral n)" by (simp only: erf_of_real [symmetric] of_real_numeral) lemma of_real_erf_1 [simp]: "of_real (erf 1) = erf 1" by (simp only: erf_of_real [symmetric] of_real_1) lemma erf_has_field_derivative: "(erf has_field_derivative of_real (2 / sqrt pi) * exp (-(z^2))) (at z within A)" proof - define a' where "a' = (\n. 2 / sqrt pi * (if even n then (- 1) ^ (n div 2) / fact (n div 2) else 0))" have "(erf has_field_derivative (\n. diffs (\n. of_real (erf_coeffs n)) n * z ^ n)) (at z)" using summable_erf unfolding erf_def by (rule termdiffs_strong_converges_everywhere) also have "diffs (\n. of_real (erf_coeffs n)) = (\n. of_real (a' n) :: 'a)" by (simp add: erf_coeffs_def a'_def diffs_def fun_eq_iff del: of_nat_Suc) hence "(\n. diffs (\n. of_real (erf_coeffs n)) n * z ^ n) = (\n. of_real (a' n) * z ^ n)" by simp also have "\ = (\n. of_real (a' (2*n)) * z ^ (2*n))" by (intro suminf_mono_reindex [symmetric]) (auto simp: strict_mono_def a'_def elim!: evenE) also have "(\n. of_real (a' (2*n)) * z ^ (2*n)) = (\n. of_real (2 / sqrt pi) * (inverse (fact n) * (-(z^2))^n))" by (simp add: fun_eq_iff power_mult [symmetric] a'_def field_simps power_minus') also have "suminf \ = of_real (2 / sqrt pi) * exp (-(z^2))" by (subst suminf_mult, intro summable_exp) (auto simp: field_simps scaleR_conv_of_real exp_def) finally show ?thesis by (rule DERIV_subset) simp_all qed lemmas erf_has_field_derivative' [derivative_intros] = erf_has_field_derivative [THEN DERIV_chain2] lemma erf_continuous_on: "continuous_on A erf" by (rule DERIV_continuous_on erf_has_field_derivative)+ lemma continuous_on_compose2_UNIV: "continuous_on UNIV g \ continuous_on s f \ continuous_on s (\x. g (f x))" by (rule continuous_on_compose2[of UNIV g s f]) simp_all lemmas erf_continuous_on' [continuous_intros] = erf_continuous_on [THEN continuous_on_compose2_UNIV] lemma erf_continuous [continuous_intros]: "continuous (at x within A) erf" by (rule continuous_within_subset[OF _ subset_UNIV]) (insert erf_continuous_on[of UNIV], auto simp: continuous_on_eq_continuous_at) lemmas erf_continuous' [continuous_intros] = continuous_within_compose2[OF _ erf_continuous] lemmas tendsto_erf [tendsto_intros] = isCont_tendsto_compose[OF erf_continuous] lemma erf_cnj [simp]: "erf (cnj z) = cnj (erf z)" proof - interpret bounded_linear cnj by (rule bounded_linear_cnj) from suminf[OF summable_erf] show ?thesis by (simp add: erf_def erf_coeffs_def) qed lemma integral_exp_minus_squared_real: assumes "a \ b" shows "((\t. exp (-(t^2))) has_integral (sqrt pi / 2 * (erf b - erf a))) {a..b}" proof - have "((\t. exp (-(t^2))) has_integral (sqrt pi / 2 * erf b - sqrt pi / 2 * erf a)) {a..b}" using assms by (intro fundamental_theorem_of_calculus) (auto intro!: derivative_eq_intros simp: has_field_derivative_iff_has_vector_derivative [symmetric]) thus ?thesis by (simp add: field_simps) qed lemma erf_real_altdef_nonneg: "x \ 0 \ erf (x::real) = 2 / sqrt pi * integral {0..x} (\t. exp (-(t^2)))" using integral_exp_minus_squared_real[of 0 x] by (simp add: has_integral_iff field_simps) lemma erf_real_altdef_nonpos: "x \ 0 \ erf (x::real) = -2 / sqrt pi * integral {0..-x} (\t. exp (-(t^2)))" using erf_real_altdef_nonneg[of "-x"] by simp lemma less_imp_erf_real_less: assumes "a < (b::real)" shows "erf a < erf b" proof - from assms have "\z. z > a \ z < b \ erf b - erf a = (b - a) * (2 / sqrt pi * exp (- z\<^sup>2))" by (intro MVT2) (auto intro!: derivative_eq_intros simp: field_simps) then guess z by (elim exE conjE) note z = this note z(3) also from assms have "(b - a) * (2 / sqrt pi * exp (- z\<^sup>2)) > 0" by (intro mult_pos_pos divide_pos_pos) simp_all finally show ?thesis by simp qed lemma le_imp_erf_real_le: "a \ (b::real) \ erf a \ erf b" by (cases "a < b") (auto dest: less_imp_erf_real_less) lemma erf_real_less_cancel [simp]: "(erf (a :: real) < erf b) \ a < b" using less_imp_erf_real_less[of a b] less_imp_erf_real_less[of b a] by (cases a b rule: linorder_cases) simp_all lemma erf_real_eq_iff [simp]: "erf (a::real) = erf b \ a = b" by (cases a b rule: linorder_cases) (auto dest: less_imp_erf_real_less) lemma erf_real_le_cancel [simp]: "(erf (a :: real) \ erf b) \ a \ b" by (cases a b rule: linorder_cases) (auto simp: less_eq_real_def) lemma inj_on_erf_real [intro]: "inj_on (erf :: real \ real) A" by (auto simp: inj_on_def) lemma strict_mono_erf_real [intro]: "strict_mono (erf :: real \ real)" by (auto simp: strict_mono_def) lemma mono_erf_real [intro]: "mono (erf :: real \ real)" by (auto simp: mono_def) lemma erf_real_ge_0_iff [simp]: "erf (x::real) \ 0 \ x \ 0" using erf_real_le_cancel[of 0 x] unfolding erf_0 . lemma erf_real_le_0_iff [simp]: "erf (x::real) \ 0 \ x \ 0" using erf_real_le_cancel[of x 0] unfolding erf_0 . lemma erf_real_gt_0_iff [simp]: "erf (x::real) > 0 \ x > 0" using erf_real_less_cancel[of 0 x] unfolding erf_0 . lemma erf_real_less_0_iff [simp]: "erf (x::real) < 0 \ x < 0" using erf_real_less_cancel[of x 0] unfolding erf_0 . lemma erf_at_top [tendsto_intros]: "((erf :: real \ real) \ 1) at_top" proof - have *: "(\n. {0..real n}) = {0..}" by (auto intro!: real_nat_ceiling_ge) let ?f = "\t::real. exp (-t\<^sup>2)" have "(\n. set_lebesgue_integral lborel {0..real n} ?f) \ set_lebesgue_integral lborel (\n. {0..real n}) ?f" using has_bochner_integral_erf_aux by (intro set_integral_cont_up ) (insert *, auto simp: incseq_def has_bochner_integral_iff set_integrable_def) also note * also have "(\n. set_lebesgue_integral lborel {0..real n} ?f) = (\n. integral {0..real n} ?f)" proof - have "\n. set_integrable lborel {0..real n} (\x. exp (- x\<^sup>2))" unfolding set_integrable_def by (intro borel_integrable_compact) (auto intro!: continuous_intros) then show ?thesis by (intro set_borel_integral_eq_integral ext) qed also have "\ = (\n. sqrt pi / 2 * erf (real n))" by (simp add: erf_real_altdef_nonneg) also have "set_lebesgue_integral lborel {0..} ?f = sqrt pi / 2" using has_bochner_integral_erf_aux by (simp add: has_bochner_integral_iff set_lebesgue_integral_def) finally have "(\n. 2 / sqrt pi * (sqrt pi / 2 * erf (real n))) \ (2 / sqrt pi) * (sqrt pi / 2)" by (intro tendsto_intros) hence "(\n. erf (real n)) \ 1" by simp thus ?thesis by (rule tendsto_sandwich_mono) auto qed lemma erf_at_bot [tendsto_intros]: "((erf :: real \ real) \ -1) at_bot" by (simp add: filterlim_at_bot_mirror tendsto_minus_cancel_left erf_at_top) lemmas tendsto_erf_at_top [tendsto_intros] = filterlim_compose[OF erf_at_top] lemmas tendsto_erf_at_bot [tendsto_intros] = filterlim_compose[OF erf_at_bot] subsection \The complimentary error function\ definition erfc where "erfc z = 1 - erf z" lemma erf_conv_erfc: "erf z = 1 - erfc z" by (simp add: erfc_def) lemma erfc_0 [simp]: "erfc 0 = 1" by (simp add: erfc_def) lemma erfc_minus: "erfc (-z) = 2 - erfc z" by (simp add: erfc_def) lemma erfc_of_real [simp]: "erfc (of_real x) = of_real (erfc x)" by (simp add: erfc_def) lemma of_real_erfc_numeral [simp]: "of_real (erfc (numeral n)) = erfc (numeral n)" by (simp add: erfc_def) lemma of_real_erfc_1 [simp]: "of_real (erfc 1) = erfc 1" by (simp add: erfc_def) lemma less_imp_erfc_real_less: "a < (b::real) \ erfc a > erfc b" by (simp add: erfc_def) lemma le_imp_erfc_real_le: "a \ (b::real) \ erfc a \ erfc b" by (simp add: erfc_def) lemma erfc_real_less_cancel [simp]: "(erfc (a :: real) < erfc b) \ a > b" by (simp add: erfc_def) lemma erfc_real_eq_iff [simp]: "erfc (a::real) = erfc b \ a = b" by (simp add: erfc_def) lemma erfc_real_le_cancel [simp]: "(erfc (a :: real) \ erfc b) \ a \ b" by (simp add: erfc_def) lemma inj_on_erfc_real [intro]: "inj_on (erfc :: real \ real) A" by (auto simp: inj_on_def) lemma antimono_erfc_real [intro]: "antimono (erfc :: real \ real)" by (auto simp: antimono_def) lemma erfc_real_ge_0_iff [simp]: "erfc (x::real) \ 1 \ x \ 0" by (simp add: erfc_def) lemma erfc_real_le_0_iff [simp]: "erfc (x::real) \ 1 \ x \ 0" by (simp add: erfc_def) lemma erfc_real_gt_0_iff [simp]: "erfc (x::real) > 1 \ x < 0" by (simp add: erfc_def) lemma erfc_real_less_0_iff [simp]: "erfc (x::real) < 1 \ x > 0" by (simp add: erfc_def) lemma erfc_has_field_derivative: "(erfc has_field_derivative -of_real (2 / sqrt pi) * exp (-(z^2))) (at z within A)" unfolding erfc_def [abs_def] by (auto intro!: derivative_eq_intros) lemmas erfc_has_field_derivative' [derivative_intros] = erfc_has_field_derivative [THEN DERIV_chain2] lemma erfc_continuous_on: "continuous_on A erfc" by (rule DERIV_continuous_on erfc_has_field_derivative)+ lemmas erfc_continuous_on' [continuous_intros] = erfc_continuous_on [THEN continuous_on_compose2_UNIV] lemma erfc_continuous [continuous_intros]: "continuous (at x within A) erfc" by (rule continuous_within_subset[OF _ subset_UNIV]) (insert erfc_continuous_on[of UNIV], auto simp: continuous_on_eq_continuous_at) lemmas erfc_continuous' [continuous_intros] = continuous_within_compose2[OF _ erfc_continuous] lemmas tendsto_erfc [tendsto_intros] = isCont_tendsto_compose[OF erfc_continuous] lemma erfc_at_top [tendsto_intros]: "((erfc :: real \ real) \ 0) at_top" unfolding erfc_def [abs_def] by (auto intro!: tendsto_eq_intros) lemma erfc_at_bot [tendsto_intros]: "((erfc :: real \ real) \ 2) at_bot" unfolding erfc_def [abs_def] by (auto intro!: tendsto_eq_intros) lemmas tendsto_erfc_at_top [tendsto_intros] = filterlim_compose[OF erfc_at_top] lemmas tendsto_erfc_at_bot [tendsto_intros] = filterlim_compose[OF erfc_at_bot] lemma integrable_exp_minus_squared: assumes "A \ {0..}" "A \ sets lborel" shows "set_integrable lborel A (\t::real. exp (-t\<^sup>2))" (is ?thesis1) and "(\t::real. exp (-t\<^sup>2)) integrable_on A" (is ?thesis2) proof - show ?thesis1 by (rule set_integrable_subset[of _ "{0..}"]) (insert assms has_bochner_integral_erf_aux, auto simp: has_bochner_integral_iff set_integrable_def) thus ?thesis2 by (rule set_borel_integral_eq_integral) qed lemma assumes "x \ 0" shows erfc_real_altdef_nonneg: "erfc x = 2 / sqrt pi * integral {x..} (\t. exp (-t\<^sup>2))" and has_integral_erfc: "((\t. exp (-t\<^sup>2)) has_integral (sqrt pi / 2 * erfc x)) {x..}" proof - let ?f = "\t::real. exp (-t\<^sup>2)" have int: "set_integrable lborel {0..} ?f" using has_bochner_integral_erf_aux by (simp add: has_bochner_integral_iff set_integrable_def) from assms have *: "{(0::real)..} = {0..x} \ {x..}" by auto have "set_lebesgue_integral lborel ({0..x} \ {x..}) ?f = set_lebesgue_integral lborel {0..x} ?f + set_lebesgue_integral lborel {x..} ?f" by (subst set_integral_Un_AE; (rule set_integrable_subset[OF int])?) (insert assms AE_lborel_singleton[of x], auto elim!: eventually_mono) also note * [symmetric] also have "set_lebesgue_integral lborel {0..} ?f = sqrt pi / 2" using has_bochner_integral_erf_aux by (simp add: has_bochner_integral_iff set_lebesgue_integral_def) also have "set_lebesgue_integral lborel {0..x} ?f = sqrt pi / 2 * erf x" by (subst set_borel_integral_eq_integral(2)[OF set_integrable_subset[OF int]]) (insert assms, auto simp: erf_real_altdef_nonneg) also have "set_lebesgue_integral lborel {x..} ?f = integral {x..} ?f" by (subst set_borel_integral_eq_integral(2)[OF set_integrable_subset[OF int]]) (insert assms, auto) finally show "erfc x = 2 / sqrt pi * integral {x..} ?f" by (simp add: field_simps erfc_def) with integrable_exp_minus_squared(2)[of "{x..}"] assms show "(?f has_integral (sqrt pi / 2 * erfc x)) {x..}" by (simp add: has_integral_iff) qed lemma erfc_real_gt_0 [simp, intro]: "erfc (x::real) > 0" proof (cases "x \ 0") case True have "0 < integral {x..x+1} (\t. exp (-(x+1)\<^sup>2))" by simp also from True have "\ \ integral {x..x+1} (\t. exp (-t\<^sup>2))" by (intro integral_le) (auto intro!: integrable_continuous_real continuous_intros power_mono) also have "\ \ sqrt pi / 2 * erfc x" by (rule has_integral_subset_le[OF _ integrable_integral has_integral_erfc]) (auto intro!: integrable_continuous_real continuous_intros True) finally have "sqrt pi / 2 * erfc x > 0" . hence "\ * (2 / sqrt pi) > 0" by (rule mult_pos_pos) simp_all thus "erfc x > 0" by simp next case False have "0 \ (1::real)" by simp also from False have "\ < erfc x" by simp finally show ?thesis . qed lemma erfc_real_less_2 [intro]: "erfc (x::real) < 2" using erfc_real_gt_0[of "-x"] unfolding erfc_minus by simp lemma erf_real_gt_neg1 [intro]: "erf (x::real) > -1" using erfc_real_less_2[of x] unfolding erfc_def by simp lemma erf_real_less_1 [intro]: "erf (x::real) < 1" using erfc_real_gt_0[of x] unfolding erfc_def by simp lemma erfc_cnj [simp]: "erfc (cnj z) = cnj (erfc z)" by (simp add: erfc_def) subsection \Specific facts about the complex case\ lemma erf_complex_altdef: "erf z = of_real (2 / sqrt pi) * contour_integral (linepath 0 z) (\t. exp (-(t^2)))" proof - define A where "A = (\z. contour_integral (linepath 0 z) (\t. exp (-(t^2))))" have [derivative_intros]: "(A has_field_derivative exp (-(z^2))) (at z)" for z :: complex unfolding A_def by (rule contour_integral_has_field_derivative[where A = UNIV]) (auto intro!: holomorphic_intros) have "erf z - 2 / sqrt pi * A z = erf 0 - 2 / sqrt pi * A 0" by (rule has_derivative_zero_unique [where f = "\z. erf z - 2 / sqrt pi * A z" and s = UNIV]) (auto intro!: has_field_derivative_imp_has_derivative derivative_eq_intros) also have "A 0 = 0" by (simp only: A_def contour_integral_trivial) finally show ?thesis unfolding A_def by (simp add: algebra_simps) qed lemma erf_holomorphic_on: "erf holomorphic_on A" by (auto simp: holomorphic_on_def field_differentiable_def intro!: erf_has_field_derivative) lemmas erf_holomorphic_on' [holomorphic_intros] = erf_holomorphic_on [THEN holomorphic_on_chain_UNIV] lemma erf_analytic_on: "erf analytic_on A" by (auto simp: analytic_on_def) (auto intro!: exI[of _ 1] holomorphic_intros) lemma erf_analytic_on' [analytic_intros]: assumes "f analytic_on A" shows "(\x. erf (f x)) analytic_on A" proof - from assms and erf_analytic_on have "erf \ f analytic_on A" by (rule analytic_on_compose_gen) auto thus ?thesis by (simp add: o_def) qed lemma erfc_holomorphic_on: "erfc holomorphic_on A" by (auto simp: holomorphic_on_def field_differentiable_def intro!: erfc_has_field_derivative) lemmas erfc_holomorphic_on' [holomorphic_intros] = erfc_holomorphic_on [THEN holomorphic_on_chain_UNIV] lemma erfc_analytic_on: "erfc analytic_on A" by (auto simp: analytic_on_def) (auto intro!: exI[of _ 1] holomorphic_intros) lemma erfc_analytic_on' [analytic_intros]: assumes "f analytic_on A" shows "(\x. erfc (f x)) analytic_on A" proof - from assms and erfc_analytic_on have "erfc \ f analytic_on A" by (rule analytic_on_compose_gen) auto thus ?thesis by (simp add: o_def) qed end diff --git a/thys/Error_Function/ROOT b/thys/Error_Function/ROOT --- a/thys/Error_Function/ROOT +++ b/thys/Error_Function/ROOT @@ -1,11 +1,11 @@ chapter AFP -session Error_Function (AFP) = "HOL-Analysis" + +session Error_Function (AFP) = "HOL-Complex_Analysis" + options [timeout = 300] sessions Landau_Symbols theories Error_Function Error_Function_Asymptotics document_files "root.tex" diff --git a/thys/Euler_MacLaurin/Euler_MacLaurin.thy b/thys/Euler_MacLaurin/Euler_MacLaurin.thy --- a/thys/Euler_MacLaurin/Euler_MacLaurin.thy +++ b/thys/Euler_MacLaurin/Euler_MacLaurin.thy @@ -1,1618 +1,1617 @@ section \The Euler--MacLaurin summation formula\ theory Euler_MacLaurin imports - "HOL-Analysis.Analysis" - "HOL-Library.Multiset" + "HOL-Complex_Analysis.Complex_Analysis" Bernoulli.Periodic_Bernpoly Bernoulli.Bernoulli_FPS begin subsection \Auxiliary facts\ (* TODO Move? *) lemma pbernpoly_of_int [simp]: "pbernpoly n (of_int a) = bernoulli n" by (simp add: pbernpoly_def) lemma continuous_on_bernpoly' [continuous_intros]: assumes "continuous_on A f" shows "continuous_on A (\x. bernpoly n (f x) :: 'a :: real_normed_algebra_1)" using continuous_on_compose2[OF continuous_on_bernpoly assms, of UNIV n] by auto lemma sum_atLeastAtMost_int_last: assumes "a < (b :: int)" shows "sum f {a..b} = sum f {a.. = sum f {a.. = f a + sum f {a<..b}" by (subst sum.insert) auto finally show ?thesis . qed lemma not_in_nonpos_Reals_imp_add_nonzero: assumes "z \ \\<^sub>\\<^sub>0" "x \ 0" shows "z + of_real x \ 0" using assms by (auto simp: add_eq_0_iff2) (* END TODO *) lemma negligible_atLeastAtMostI: "b \ a \ negligible {a..(b::real)}" by (cases "b < a") auto lemma integrable_on_negligible: "negligible A \ (f :: 'n :: euclidean_space \ 'a :: banach) integrable_on A" by (subst integrable_spike_set_eq[of _ "{}"]) (simp_all add: integrable_on_empty) lemma Union_atLeastAtMost_real_of_int: assumes "a < b" shows "(\n\{a.. {real_of_int a..real_of_int b}" thus "x \ (\n\{a.. real_of_int a" "x < real_of_int b" by simp_all hence "x \ of_int \x\" "x \ of_int \x\ + 1" by linarith+ moreover from x have "\x\ \ a" "\x\ < b" by linarith+ ultimately have "\n\{a.. {of_int n..of_int (n + 1)}" by (intro bexI[of _ "\x\"]) simp_all thus ?thesis by blast qed qed auto subsection \The remainder terms\ text \ The following describes the remainder term in the classical version of the Euler--MacLaurin formula. \ definition EM_remainder' :: "nat \ (real \ 'a :: banach) \ real \ real \ 'a" where "EM_remainder' n f a b = ((-1) ^ Suc n / fact n) *\<^sub>R integral {a..b} (\t. pbernpoly n t *\<^sub>R f t)" text \ Next, we define the remainder term that occurs when one lets the right bound of summation in the Euler--MacLaurin formula tend to infinity. \ definition EM_remainder_converges :: "nat \ (real \ 'a :: banach) \ int \ bool" where "EM_remainder_converges n f a \ (\L. ((\x. EM_remainder' n f a (of_int x)) \ L) at_top)" definition EM_remainder :: "nat \ (real \ 'a :: banach) \ int \ 'a" where "EM_remainder n f a = (if EM_remainder_converges n f a then Lim at_top (\x. EM_remainder' n f a (of_int x)) else 0)" text \ The following lemmas are fairly obvious -- but tedious to prove -- properties of the remainder terms. \ lemma EM_remainder_eqI: fixes L assumes "((\x. EM_remainder' n f b (of_int x)) \ L) at_top" shows "EM_remainder n f b = L" using assms by (auto simp: EM_remainder_def EM_remainder_converges_def intro!: tendsto_Lim) lemma integrable_EM_remainder'_int: fixes a b :: int and f :: "real \ 'a :: banach" assumes "continuous_on {of_int a..of_int b} f" shows "(\t. pbernpoly n t *\<^sub>R f t) integrable_on {a..b}" proof - have [continuous_intros]: "continuous_on A f" if "A \ {of_int a..of_int b}" for A using continuous_on_subset[OF assms that] . consider "a > b" | "a = b" | "a < b" "n = 1" | "a < b" "n \ 1" by (cases a b rule: linorder_cases) auto thus ?thesis proof cases assume "a < b" and "n \ 1" thus ?thesis by (intro integrable_continuous_real continuous_intros) auto next assume ab: "a < b" and [simp]: "n = 1" let ?A = "(\n. {real_of_int n..real_of_int (n+1)}) ` {a.. ?A" then obtain i where i: "i \ {a..t. pbernpoly n t *\<^sub>R f t) integrable_on I" proof (rule integrable_spike) show "(\t. (t - of_int i - 1/2) *\<^sub>R f t) integrable_on I" using i by (auto intro!: integrable_continuous_real continuous_intros) next fix x assume "x \ I - {of_int (i + 1)}" with i have "of_int i \ x" "x < of_int i + 1" by simp_all hence "floor x = i" by linarith thus "pbernpoly n x *\<^sub>R f x = (x - of_int i - 1 / 2) *\<^sub>R f x" by (simp add: pbernpoly_def bernpoly_def frac_def) qed simp_all qed qed (simp_all add: integrable_on_negligible) qed lemma integrable_EM_remainder': fixes a b :: real and f :: "real \ 'a :: banach" assumes "continuous_on {a..b} f" shows "(\t. pbernpoly n t *\<^sub>R f t) integrable_on {a..b}" proof (cases "\a\ \ \b\") case True define a' b' where "a' = \a\" and "b' = \b\" from True have *: "a' \ b'" "a' \ a" "b' \ b" by (auto simp: a'_def b'_def) from * have A: "(\t. pbernpoly n t *\<^sub>R f t) integrable_on ({a'..b'})" by (intro integrable_EM_remainder'_int continuous_on_subset[OF assms]) auto have B: "(\t. pbernpoly n t *\<^sub>R f t) integrable_on ({a..a'})" proof (rule integrable_spike) show "pbernpoly n x *\<^sub>R f x = bernpoly n (x - of_int (floor a)) *\<^sub>R f x" if "x \ {a..real_of_int a'} - {real_of_int a'}"for x proof - have "x \ a" "x t. pbernpoly n t *\<^sub>R f t) integrable_on ({b'..b})" proof (rule integrable_spike) show "pbernpoly n x *\<^sub>R f x = bernpoly n (x - of_int b') *\<^sub>R f x" if "x \ {real_of_int b'..b} - {real_of_int b'}" for x proof - have "x \ b" "x > real_of_int b'" using that by auto with True have "floor x = b'" unfolding b'_def by (intro floor_unique; linarith) thus ?thesis by (simp add: pbernpoly_def frac_def) qed qed (insert *, auto intro!: continuous_intros integrable_continuous_real continuous_on_subset[OF assms]) have "(\t. pbernpoly n t *\<^sub>R f t) integrable_on ({a..a'} \ {a'..b'} \ {b'..b})" using * A B C by (intro integrable_Un; (subst ivl_disj_un)?) (auto simp: ivl_disj_un max_def min_def) also have "{a..a'} \ {a'..b'} \ {b'..b} = {a..b}" using * by auto finally show ?thesis . next assume *: "\ceiling a \ floor b" show ?thesis proof (rule integrable_spike) show "(\t. bernpoly n (t - floor a) *\<^sub>R f t) integrable_on {a..b}" using * by (auto intro!: integrable_continuous_real continuous_intros assms) next show "pbernpoly n x *\<^sub>R f x = bernpoly n (x - floor a) *\<^sub>R f x" if "x \ {a..b} - {}" for x proof - from * have **: "b < floor a + 1" unfolding ceiling_altdef by (auto split: if_splits simp: le_floor_iff) from that have x: "x \ a" "x \ b" by simp_all with * ** have "floor x = floor a" by linarith thus ?thesis by (simp add: pbernpoly_def frac_def) qed qed simp_all qed lemma EM_remainder'_bounded_linear: assumes "bounded_linear h" assumes "continuous_on {a..b} f" shows "EM_remainder' n (\x. h (f x)) a b = h (EM_remainder' n f a b)" proof - have "integral {a..b} (\t. pbernpoly n t *\<^sub>R h (f t)) = integral {a..b} (\t. h (pbernpoly n t *\<^sub>R f t))" using assms by (simp add: linear_simps) also have "\ = h (integral {a..b} (\t. pbernpoly n t *\<^sub>R f t))" by (subst integral_linear [OF _ assms(1), symmetric]) (auto intro!: integrable_EM_remainder' assms(2) simp: o_def) finally show ?thesis using assms(1) by (simp add: EM_remainder'_def linear_simps) qed lemma EM_remainder_converges_of_real: assumes "EM_remainder_converges n f a" "continuous_on {of_int a..} f" shows "EM_remainder_converges n (\x. of_real (f x)) a" proof - from assms obtain L where L: "((\b. EM_remainder' n f (real_of_int a) (real_of_int b)) \ L) at_top" by (auto simp: EM_remainder_converges_def) have "((\b. EM_remainder' n (\x. of_real (f x)) (of_int a) (of_int b)) \ of_real L) at_top" proof (rule Lim_transform_eventually) show "eventually (\b. of_real (EM_remainder' n f (of_int a) (of_int b)) = EM_remainder' n (\x. of_real (f x)) (of_int a) (of_int b)) at_top" using eventually_ge_at_top[of a] by eventually_elim (intro EM_remainder'_bounded_linear [OF bounded_linear_of_real, symmetric] continuous_on_subset[OF assms(2)], auto) qed (intro tendsto_intros L) thus ?thesis unfolding EM_remainder_converges_def .. qed lemma EM_remainder_converges_of_real_iff: fixes f :: "real \ real" assumes "continuous_on {of_int a..} f" shows "EM_remainder_converges n (\x. of_real (f x) :: 'a :: {banach, real_normed_algebra_1, real_inner}) a \ EM_remainder_converges n f a" proof assume "EM_remainder_converges n (\x. of_real (f x) :: 'a) a" then obtain L :: 'a where L: "((\b. EM_remainder' n (\x. of_real (f x)) (of_int a) (of_int b)) \ L) at_top" by (auto simp: EM_remainder_converges_def) have "((\b. EM_remainder' n f (of_int a) (of_int b)) \ L \ 1) at_top" proof (rule Lim_transform_eventually) show "eventually (\b. EM_remainder' n (\x. of_real (f x) :: 'a) (of_int a) (of_int b) \ 1 = EM_remainder' n f (of_int a) (of_int b)) at_top" using eventually_ge_at_top[of a] by eventually_elim (subst EM_remainder'_bounded_linear [OF bounded_linear_of_real], auto intro!: continuous_on_subset[OF assms]) qed (intro tendsto_intros L) thus "EM_remainder_converges n f a" unfolding EM_remainder_converges_def .. qed (intro EM_remainder_converges_of_real assms) lemma EM_remainder_of_real: assumes "continuous_on {a..} f" shows "EM_remainder n (\x. of_real (f x) :: 'a :: {banach, real_normed_algebra_1, real_inner}) a = of_real (EM_remainder n f a)" proof - have eq: "EM_remainder' n (\x. of_real (f x) :: 'a) (real_of_int a) = (\x::int. of_real (EM_remainder' n f a x))" by (intro ext EM_remainder'_bounded_linear[OF bounded_linear_of_real] continuous_on_subset[OF assms]) auto show ?thesis proof (cases "EM_remainder_converges n f a") case False with EM_remainder_converges_of_real_iff[OF assms, of n] show ?thesis by (auto simp: EM_remainder_def) next case True then obtain L where L: "((\x. EM_remainder' n f a (real_of_int x)) \ L) at_top" by (auto simp: EM_remainder_converges_def) have L': "((\x. EM_remainder' n (\x. of_real (f x) :: 'a) a (real_of_int x)) \ of_real L) at_top" unfolding eq by (intro tendsto_of_real L) from L L' tendsto_Lim[OF _ L] tendsto_Lim[OF _ L'] show ?thesis by (auto simp: EM_remainder_def EM_remainder_converges_def) qed qed lemma EM_remainder'_cong: assumes "\x. x \ {a..b} \ f x = g x" "n = n'" "a = a'" "b = b'" shows "EM_remainder' n f a b = EM_remainder' n' g a' b'" proof - have "integral {a..b} (\t. pbernpoly n t *\<^sub>R f t) = integral {a'..b'} (\t. pbernpoly n' t *\<^sub>R g t)" unfolding assms using assms by (intro integral_cong) auto with assms show ?thesis by (simp add: EM_remainder'_def) qed lemma EM_remainder_converges_cong: assumes "\x. x \ of_int a \ f x = g x" "n = n'" "a = a'" shows "EM_remainder_converges n f a = EM_remainder_converges n' g a'" unfolding EM_remainder_converges_def by (subst EM_remainder'_cong[OF _ refl refl refl, of _ _ f g]) (use assms in auto) lemma EM_remainder_cong: assumes "\x. x \ of_int a \ f x = g x" "n = n'" "a = a'" shows "EM_remainder n f a = EM_remainder n' g a'" proof - have *: "EM_remainder_converges n f a = EM_remainder_converges n' g a'" using assms by (intro EM_remainder_converges_cong) auto show ?thesis unfolding EM_remainder_def by (subst EM_remainder'_cong[OF _ refl refl refl, of _ _ f g]) (use assms * in auto) qed lemma EM_remainder_converges_cnj: assumes "continuous_on {a..} f" and "EM_remainder_converges n f a" shows "EM_remainder_converges n (\x. cnj (f x)) a" proof - interpret bounded_linear cnj by (rule bounded_linear_cnj) obtain L where L: "((\x. EM_remainder' n f (real_of_int a) (real_of_int x)) \ L) at_top" using assms unfolding EM_remainder_converges_def by blast note tendsto_cnj [OF this] also have "(\x. cnj (EM_remainder' n f (real_of_int a) (real_of_int x))) = (\x. EM_remainder' n (\x. cnj (f x)) (real_of_int a) (real_of_int x))" by (subst EM_remainder'_bounded_linear [OF bounded_linear_cnj]) (rule continuous_on_subset [OF assms(1)], auto) finally have L': "(\ \ cnj L) at_top" . thus "EM_remainder_converges n (\x. cnj (f x)) a" by (auto simp: EM_remainder_converges_def) qed lemma EM_remainder_converges_cnj_iff: assumes "continuous_on {of_int a..} f" shows "EM_remainder_converges n (\x. cnj (f x)) a \ EM_remainder_converges n f a" proof assume "EM_remainder_converges n (\x. cnj (f x)) a" hence "EM_remainder_converges n (\x. cnj (cnj (f x))) a" by (rule EM_remainder_converges_cnj [rotated]) (auto intro: continuous_intros assms) thus "EM_remainder_converges n f a" by simp qed (intro EM_remainder_converges_cnj assms) lemma EM_remainder_cnj: assumes "continuous_on {a..} f" shows "EM_remainder n (\x. cnj (f x)) a = cnj (EM_remainder n f a)" proof (cases "EM_remainder_converges n f a") case False hence "\EM_remainder_converges n (\x. cnj (f x)) a" by (subst EM_remainder_converges_cnj_iff [OF assms]) with False show ?thesis by (simp add: EM_remainder_def) next case True then obtain L where L: "((\x. EM_remainder' n f (real_of_int a) (real_of_int x)) \ L) at_top" unfolding EM_remainder_converges_def by blast note tendsto_cnj [OF this] also have "(\x. cnj (EM_remainder' n f (real_of_int a) (real_of_int x))) = (\x. EM_remainder' n (\x. cnj (f x)) (real_of_int a) (real_of_int x))" by (subst EM_remainder'_bounded_linear [OF bounded_linear_cnj]) (rule continuous_on_subset [OF assms(1)], auto) finally have L': "(\ \ cnj L) at_top" . moreover from assms and L have "EM_remainder n f a = L" by (intro EM_remainder_eqI) ultimately show "EM_remainder n (\x. cnj (f x)) a = cnj (EM_remainder n f a)" using L' by (intro EM_remainder_eqI) simp_all qed lemma EM_remainder'_combine: fixes f :: "real \ 'a :: banach" assumes [continuous_intros]: "continuous_on {a..c} f" assumes "a \ b" "b \ c" shows "EM_remainder' n f a b + EM_remainder' n f b c = EM_remainder' n f a c" proof - have "integral {a..b} (\t. pbernpoly n t *\<^sub>R f t) + integral {b..c} (\t. pbernpoly n t *\<^sub>R f t) = integral {a..c} (\t. pbernpoly n t *\<^sub>R f t)" by (intro Henstock_Kurzweil_Integration.integral_combine assms integrable_EM_remainder') from this [symmetric] show ?thesis by (simp add: EM_remainder'_def algebra_simps) qed lemma uniformly_convergent_EM_remainder': fixes f :: "'a \ real \ 'b :: {banach,real_normed_algebra}" assumes deriv: "\y. a \ y \ (G has_real_derivative g y) (at y within {a..})" assumes integrable: "\a' b y. y \ A \ a \ a' \ a' \ b \ (\t. pbernpoly n t *\<^sub>R f y t) integrable_on {a'..b}" assumes conv: "convergent (\y. G (real y))" assumes bound: "eventually (\x. \y\A. norm (f y x) \ g x) at_top" shows "uniformly_convergent_on A (\b s. EM_remainder' n (f s) a b)" proof - interpret bounded_linear "\x::'b. ((- 1) ^ Suc n / fact n) *\<^sub>R x" by (rule bounded_linear_scaleR_right) from bounded_pbernpoly[of n] guess C . note C = this from C[of 0] have [simp]: "C \ 0" by simp show ?thesis unfolding EM_remainder'_def proof (intro uniformly_convergent_on uniformly_convergent_improper_integral') fix x assume "x \ a" thus "((\x. C * G x) has_real_derivative C * g x) (at x within {a..})" by (intro DERIV_cmult deriv) next fix y a' b assume "y \ A" "a \ a'" "a' \ b" thus "(\t. pbernpoly n t *\<^sub>R f y t) integrable_on {a'..b}" by (rule integrable) next from conv obtain L where "(\y. G (real y)) \ L" by (auto simp: convergent_def) from tendsto_mult[OF tendsto_const[of C] this] show "convergent (\y. C * G (real y))" by (auto simp: convergent_def) next show "\\<^sub>F x in at_top. \y\A. norm (pbernpoly n x *\<^sub>R f y x) \ C * g x" using C unfolding norm_scaleR by (intro eventually_mono[OF bound] ballI mult_mono) auto qed qed lemma uniform_limit_EM_remainder: fixes f :: "'a \ real \ 'b :: {banach,real_normed_algebra}" assumes deriv: "\y. a \ y \ (G has_real_derivative g y) (at y within {a..})" assumes integrable: "\a' b y. y \ A \ a \ a' \ a' \ b \ (\t. pbernpoly n t *\<^sub>R f y t) integrable_on {a'..b}" assumes conv: "convergent (\y. G (real y))" assumes bound: "eventually (\x. \y\A. norm (f y x) \ g x) at_top" shows "uniform_limit A (\b s. EM_remainder' n (f s) a b) (\s. EM_remainder n (f s) a) sequentially" proof - have *: "uniformly_convergent_on A (\b s. EM_remainder' n (f s) a b)" by (rule uniformly_convergent_EM_remainder'[OF assms]) also have "?this \ ?thesis" unfolding uniformly_convergent_uniform_limit_iff proof (intro uniform_limit_cong refl always_eventually allI ballI) fix s assume "s \ A" with * have **: "convergent (\b. EM_remainder' n (f s) a b)" by (rule uniformly_convergent_imp_convergent) show "lim (\b. EM_remainder' n (f s) a b) = EM_remainder n (f s) a" proof (rule sym, rule EM_remainder_eqI) have "((\x. EM_remainder' n (f s) (real_of_int a) (real x)) \ lim (\x. EM_remainder' n (f s) (real_of_int a) (real x))) at_top" (is "(_ \ ?L) _") using ** unfolding convergent_LIMSEQ_iff by blast hence "((\x. EM_remainder' n (f s) (real_of_int a) (real (nat x))) \ ?L) at_top" by (rule filterlim_compose) (fact filterlim_nat_sequentially) thus "((\x. EM_remainder' n (f s) (real_of_int a) (real_of_int x)) \ ?L) at_top" by (rule Lim_transform_eventually) (auto intro: eventually_mono[OF eventually_ge_at_top[of 0]]) qed qed finally show \ . qed lemma tendsto_EM_remainder: fixes f :: "real \ 'b :: {banach,real_normed_algebra}" assumes deriv: "\y. a \ y \ (G has_real_derivative g y) (at y within {a..})" assumes integrable: "\a' b . a \ a' \ a' \ b \ (\t. pbernpoly n t *\<^sub>R f t) integrable_on {a'..b}" assumes conv: "convergent (\y. G (real y))" assumes bound: "eventually (\x. norm (f x) \ g x) at_top" shows "filterlim (\b. EM_remainder' n f a b) (nhds (EM_remainder n f a)) sequentially" proof - have "uniform_limit {()} (\b s. EM_remainder' n f a b) (\s. EM_remainder n f a) sequentially" using assms by (intro uniform_limit_EM_remainder[where G = G and g = g]) auto moreover have "() \ {()}" by simp ultimately show ?thesis by (rule tendsto_uniform_limitI) qed lemma EM_remainder_0 [simp]: "EM_remainder n (\x. 0) a = 0" by (rule EM_remainder_eqI) (simp add: EM_remainder'_def) lemma holomorphic_EM_remainder': assumes deriv: "\z t. z \ U \ t \ {a..x} \ ((\z. f z t) has_field_derivative f' z t) (at z within U)" assumes int: "\b c z e. a \ b \ c \ x \ z \ U \ (\t. of_real (bernpoly n (t - e)) * f z t) integrable_on {b..c}" assumes cont: "continuous_on (U \ {a..x}) (\(z, t). f' z t)" assumes "convex U" shows "(\s. EM_remainder' n (f s) a x) holomorphic_on U" unfolding EM_remainder'_def scaleR_conv_of_real proof (intro holomorphic_intros) have holo: "(\z. integral (cbox b c) (\t. of_real (bernpoly n (t - e)) * f z t)) holomorphic_on U" if "b \ a" "c \ x" for b c e :: real proof (rule leibniz_rule_holomorphic) fix z t assume "z \ U" "t \ cbox b c" thus "((\z. complex_of_real (bernpoly n (t - e)) * f z t) has_field_derivative complex_of_real (bernpoly n (t - e)) * f' z t) (at z within U)" using that by (intro DERIV_cmult deriv) auto next fix z assume "z \ U" thus "(\t. complex_of_real (bernpoly n (t - e)) * f z t) integrable_on cbox b c" using that int[of b c z] by auto next have "continuous_on (U \ {b..c}) (\(z, t). f' z t)" using cont by (rule continuous_on_subset) (insert that, auto) thus "continuous_on (U \ cbox b c) (\(z, t). complex_of_real (bernpoly n (t - e)) * f' z t)" by (auto simp: case_prod_unfold intro!: continuous_intros) qed fact+ consider "a > x" | "a \ x" "floor x \ a" | "a \ x" "floor x > a" by force hence "(\z. integral (cbox a x) (\t. of_real (pbernpoly n t) * f z t)) holomorphic_on U" (is "?f a x holomorphic_on _") proof cases case 2 have "(\z. integral (cbox a x) (\t. of_real (bernpoly n (t - of_int \x\)) * f z t)) holomorphic_on U" by (intro holo) auto also have "(\z. integral (cbox a x) (\t. of_real (bernpoly n (t - of_int \x\)) * f z t)) = ?f a x" proof (intro ext integral_cong, goal_cases) case (1 z t) hence "t \ a" "t \ x" by auto hence "floor t = floor x" using 2 by linarith thus ?case by (simp add: pbernpoly_def frac_def) qed finally show ?thesis . next case 3 define N :: "int set" where "N = {\a\..<\x\}" define A where "A = insert {a..of_int \a\} (insert {of_int \x\..x} ((\n. {of_int n..of_int n + 1}) ` N))" { fix X assume "X \ A" then consider "X = {a..of_int \a\}" | "X = {of_int \x\..x}" | n where "X = {of_int n..of_int n + 1}" "n \ N" by (auto simp: A_def) } note A_cases = this have division: "A division_of {a..x}" proof (rule division_ofI) show "finite A" by (auto simp: A_def N_def) fix K assume K: "K \ A" from 3 have "of_int \a\ \ x" using ceiling_le[of a "floor x"] by linarith moreover from 3 have "of_int \x\ \ a" by linarith ultimately show "K \ {a..x}" using K 3 by (auto simp: A_def N_def) linarith+ from K show "K \ {}" and "\a b. K = cbox a b" by (auto simp: A_def) next fix K1 K2 assume K: "K1 \ A" "K2 \ A" "K1 \ K2" have F1: "interior {a..\a\} \ interior {\x\..x} = {}" using 3 ceiling_le[of a "floor x"] by (auto simp: min_def max_def) hence F2: "interior {\x\..x} \ interior {a..\a\} = {}" by simp have F3: "interior {a..\a\} \ interior {of_int n..of_int n+1} = {}" "interior {\x\..x} \ interior {of_int n..of_int n+1} = {}" "interior {of_int n..of_int n+1} \ interior {a..\a\} = {}" "interior {of_int n..of_int n+1} \ interior {\x\..x} = {}"if "n \ N" for n using 3 ceiling_le[of a "floor x"] that by (auto simp: min_def max_def N_def) have F4: "interior {real_of_int n..of_int n+1} \ interior {of_int m..of_int m+1} = {}" if "{real_of_int n..of_int n+1} \ {of_int m..of_int m+1}" for m n proof - from that have "n \ m" by auto thus ?thesis by simp qed from F1 F2 F3 F4 K show "interior K1 \ interior K2 = {}" by (elim A_cases) (simp_all only: not_False_eq_True) next show "\A = {a..x}" proof (cases "\a\ = \x\") case True thus ?thesis using 3 by (auto simp: A_def N_def intro: order.trans) linarith+ next case False with 3 have *: "\a\ < \x\" by linarith have "\A = {a..of_int \a\} \ (\n\N. {of_int n..of_int (n + 1)}) \ {of_int \x\..x}" by (simp add: A_def Un_ac) also have "(\n\N. {of_int n..of_int (n + 1)}) = {of_int \a\..real_of_int \x\}" using * unfolding N_def by (intro Union_atLeastAtMost_real_of_int) also have "{a..of_int \a\} \ \ = {a..real_of_int \x\}" using 3 * by (intro ivl_disj_un) auto also have "\ \ {of_int \x\..x} = {a..x}" using 3 * by (intro ivl_disj_un) auto finally show ?thesis . qed qed have "(\z. \X\A. integral X (\t. of_real (bernpoly n (t - \Inf X\)) * f z t)) holomorphic_on U" proof (intro holomorphic_on_sum holo, goal_cases) case (1 X) from 1 and division have subset: "X \ {a..x}" by (auto simp: division_of_def) from 1 obtain b c where [simp]: "X = cbox b c" "b \ c" by (auto simp: A_def) from subset have "b \ a" "c \ x" by auto hence "(\x. integral (cbox b c) (\t. of_real (bernpoly n (t - \Inf {b..c}\)) * f x t)) holomorphic_on U" by (intro holo) auto thus ?case by simp qed also have "?this \ (\z. integral {a..x} (\t. of_real (pbernpoly n t) * f z t)) holomorphic_on U" proof (intro holomorphic_cong refl, goal_cases) case (1 z) have "((\t. of_real (pbernpoly n t) * f z t) has_integral (\X\A. integral X (\t. of_real (bernpoly n (t - \Inf X\)) * f z t))) {a..x}" using division proof (rule has_integral_combine_division) fix X assume X: "X \ A" then obtain b c where X': "X = {b..c}" "b \ c" by (elim A_cases) auto from X and division have "X \ {a..x}" by (auto simp: division_of_def) with X' have bc: "b \ a" "c \ x" by auto have "((\t. of_real (bernpoly n (t - of_int \Inf X\)) * f z t) has_integral integral X (\t. of_real (bernpoly n (t - of_int \Inf X\)) * f z t)) X" unfolding X' using \z \ U\ bc by (intro integrable_integral int) also have "?this \ ((\t. of_real (pbernpoly n t) * f z t) has_integral integral X (\t. of_real (bernpoly n (t - of_int \Inf X\)) * f z t)) X" proof (rule has_integral_spike_eq[of "{Sup X}"], goal_cases) case (2 t) note t = this from \X \ A\ have "\t\ = \Inf X\" proof (cases rule: A_cases [consumes 1]) case 1 with t show ?thesis by (intro floor_unique) (auto simp: ceiling_altdef split: if_splits, (linarith+)?) next case 2 with t show ?thesis by (intro floor_unique) (auto simp: ceiling_altdef split: if_splits, (linarith+)?) next case 3 with t show ?thesis by (intro floor_unique) (auto simp: ceiling_altdef N_def split: if_splits) qed thus ?case by (simp add: pbernpoly_def frac_def) qed auto finally show \ . qed thus ?case by (simp add: has_integral_iff) qed finally show ?thesis by simp qed auto thus "(\z. integral {a..x} (\t. of_real (pbernpoly n t) * f z t)) holomorphic_on U" by simp qed lemma assumes deriv: "\y. a \ y \ (G has_real_derivative g y) (at y within {a..})" assumes deriv': "\z t x. z \ U \ x \ a \ t \ {a..x} \ ((\z. f z t) has_field_derivative f' z t) (at z within U)" assumes cont: "continuous_on (U \ {of_int a..}) (\(z, t). f' z t)" assumes int: "\b c z e. a \ b \ z \ U \ (\t. of_real (bernpoly n (t - e)) * f z t) integrable_on {b..c}" assumes int': "\a' b y. y \ U \ a \ a' \ a' \ b \ (\t. pbernpoly n t *\<^sub>R f y t) integrable_on {a'..b}" assumes conv: "convergent (\y. G (real y))" assumes bound: "eventually (\x. \y\U. norm (f y x) \ g x) at_top" assumes "open U" shows analytic_EM_remainder: "(\s::complex. EM_remainder n (f s) a) analytic_on U" and holomorphic_EM_remainder: "(\s::complex. EM_remainder n (f s) a) holomorphic_on U" proof - show "(\s::complex. EM_remainder n (f s) a) analytic_on U" unfolding analytic_on_def proof fix z assume "z \ U" from \z \ U\ and \open U\ obtain \ where \: "\ > 0" "ball z \ \ U" by (auto simp: open_contains_ball) have "(\s. EM_remainder n (f s) a) holomorphic_on ball z \" proof (rule holomorphic_uniform_sequence) fix x :: nat show "(\s. EM_remainder' n (f s) a x) holomorphic_on ball z \" proof (rule holomorphic_EM_remainder', goal_cases) fix s t assume "s \ ball z \" "t \ {real_of_int a..real x}" thus "((\z. f z t) has_field_derivative f' s t) (at s within ball z \)" using \ by (intro DERIV_subset[OF deriv'[of _ x]]) auto next case (2 b c s e) with \ have "s \ U" by blast with 2 show ?case using \ int[of b s e c] by (cases "a \ x") auto next from cont show "continuous_on (ball z \ \ {real_of_int a..real x}) (\(z, t). f' z t)" by (rule continuous_on_subset) (insert \, auto) qed (auto) next fix s assume s: "s \ ball z \" have "open (ball z \)" by simp with s obtain \ where \: "\ > 0" "cball s \ \ ball z \" unfolding open_contains_cball by blast moreover have bound': "eventually (\x. \y\cball s \. norm (f y x) \ g x) at_top" by (intro eventually_mono [OF bound]) (insert \ \, auto) have "uniform_limit (cball s \) (\x s. EM_remainder' n (f s) (real_of_int a) (real x)) (\s. EM_remainder n (f s) a) sequentially" by (rule uniform_limit_EM_remainder[OF deriv int' conv bound']) (insert \ \ s, auto) ultimately show "\\>0. cball s \ \ ball z \ \ uniform_limit (cball s \) (\x s. EM_remainder' n (f s) (real_of_int a) (real x)) (\s. EM_remainder n (f s) a) sequentially" by blast qed auto with \ show "\\>0. (\s. EM_remainder n (f s) a) holomorphic_on ball z \" by blast qed thus "(\s::complex. EM_remainder n (f s) a) holomorphic_on U" by (rule analytic_imp_holomorphic) qed text \ The following lemma is the first step in the proof of the Euler--MacLaurin formula: We show the relationship between the first-order remainder term and the difference of the integral and the sum. \ context fixes f f' :: "real \ 'a :: banach" fixes a b :: int and I S :: 'a fixes Y :: "real set" assumes "a \ b" assumes fin: "finite Y" assumes cont: "continuous_on {real_of_int a..real_of_int b} f" assumes deriv [derivative_intros]: "\x::real. x \ {a..b} - Y \ (f has_vector_derivative f' x) (at x)" defines S_def: "S \ (\i\{a<..b}. f i)" and I_def: "I \ integral {a..b} f" begin lemma diff_sum_integral_has_integral_int: "((\t. (frac t - 1/2) *\<^sub>R f' t) has_integral (S - I - (f b - f a) /\<^sub>R 2)) {a..b}" proof (cases "a = b") case True thus ?thesis by (simp_all add: S_def I_def has_integral_refl) next case False with \a \ b\ have ab: "a < b" by simp let ?A = "(\n. {real_of_int n..real_of_int (n+1)}) ` {a.. {of_int a..of_int b}" for A using continuous_on_subset[OF cont that] . define d where "d = (\x. (f x + f (x + 1)) /\<^sub>R 2 - integral {x..x+1} f)" have "((\t. (frac t - 1/2) *\<^sub>R f' t) has_integral d i) {of_int i..of_int (i+1)}" if i: "i \ {a..R f' x = (x - of_int i - 1 / 2) *\<^sub>R f' x" if "x \ {of_int i..of_int (i + 1)} - {of_int (i + 1)}" for x proof - have "x \ of_int i" "x < of_int (i + 1)" using that by auto hence "floor x = of_int i" by (subst floor_unique) auto thus ?thesis by (simp add: frac_def) qed next define h where "h = (\x::real. (x - of_int i - 1 / 2) *\<^sub>R f' x)" define g where "g = (\x::real. (x - of_int i - 1/2) *\<^sub>R f x - integral {of_int i..x} f)" have *: "((\x. integral {real_of_int i..x} f) has_vector_derivative f x) (at x within {i..i+1})" if "x \ {of_int i<..x. integral {real_of_int i..x} f) has_vector_derivative f x) (at x)" if "x \ {of_int i<..I. I\?A \ ((\x. (frac x - 1 / 2) *\<^sub>R f' x) has_integral d (\Inf I\)) I" by (auto simp: add_ac) have "((\x::real. (frac x - 1 / 2) *\<^sub>R f' x) has_integral (\I\?A. d (\Inf I\))) (\?A)" by (intro has_integral_Union * finite_imageI) (force intro!: negligible_atLeastAtMostI pairwiseI)+ also have "\?A = {of_int a..of_int b}" by (intro Union_atLeastAtMost_real_of_int ab) also have "(\I\?A. d (\Inf I\)) = (\i=a.. = (1 / 2) *\<^sub>R ((\i = a..i = a..i = a..R (?S1 + ?S2) - ?S3") by (simp add: d_def algebra_simps sum.distrib sum_subtractf scaleR_sum_right) also have "?S1 = (\i = a..b. f (real_of_int i)) - f b" unfolding S_def using ab by (subst sum_atLeastAtMost_int_last) auto also have "(\i = a..b. f (real_of_int i)) = S + f a" unfolding S_def using ab by (subst sum_atLeastAtMost_int_head) auto also have "?S2 = S" unfolding S_def by (intro sum.reindex_bij_witness[of _ "\i. i-1" "\i. i+1"]) auto also have "(1 / 2) *\<^sub>R (S + f a - f b + S) = (1/2) *\<^sub>R S + (1/2) *\<^sub>R S - (f b - f a) /\<^sub>R 2" by (simp add: algebra_simps) also have "(1/2) *\<^sub>R S + (1/2) *\<^sub>R S = S" by (simp add: scaleR_add_right [symmetric]) also have "?S3 = (\I\?A. integral I f)" by (subst sum.reindex) (auto simp: inj_on_def add_ac) also have "\ = I" unfolding I_def by (intro integral_combine_division_topdown [symmetric] division integrable_continuous_real continuous_intros) simp_all finally show ?thesis by (simp add: algebra_simps) qed lemma diff_sum_integral_has_integral_int': "((\t. pbernpoly 1 t *\<^sub>R f' t) has_integral (S - I - (f b - f a) /\<^sub>R 2 )) {a..b}" using diff_sum_integral_has_integral_int by (simp add: pbernpoly_def bernpoly_def) lemma EM_remainder'_Suc_0: "EM_remainder' (Suc 0) f' a b = S - I - (f b - f a) /\<^sub>R 2" using diff_sum_integral_has_integral_int' by (simp add: has_integral_iff EM_remainder'_def) end text \ Next, we show that the $n$-th-order remainder can be expressed in terms of the $n+1$-th-order remainder term. Iterating this essentially yields the Euler--MacLaurin formula. \ context fixes f f' :: "real \ 'a :: banach" and a b :: int and n :: nat and A :: "real set" assumes ab: "a \ b" and n: "n > 0" assumes fin: "finite A" assumes cont: "continuous_on {of_int a..of_int b} f" assumes cont': "continuous_on {of_int a..of_int b} f'" assumes deriv: "\x. x \ {of_int a<.. (f has_vector_derivative f' x) (at x)" begin lemma EM_remainder'_integral_conv_Suc: shows "integral {a..b} (\t. pbernpoly n t *\<^sub>R f t) = (bernoulli (Suc n) / real (Suc n)) *\<^sub>R (f b - f a) - integral {a..b} (\t. pbernpoly (Suc n) t *\<^sub>R f' t) /\<^sub>R real (Suc n)" unfolding EM_remainder'_def proof - let ?h = "\i. (pbernpoly (Suc n) (real_of_int i) / real (Suc n)) *\<^sub>R f (real_of_int i)" define T where "T = integral {a..b} (\t. (pbernpoly (Suc n) t / real (Suc n)) *\<^sub>R f' t)" note [derivative_intros] = has_field_derivative_pbernpoly_Suc' let ?A = "real_of_int ` {a..b} \ A" have "((\t. pbernpoly n t *\<^sub>R f t) has_integral (-T + (?h b - ?h a))) {a..b}" proof (rule integration_by_parts_interior_strong[OF bounded_bilinear_scaleR]) from fin show "finite ?A" by simp from \n > 0\ show "continuous_on {of_int a..of_int b} (\t. pbernpoly (Suc n) t / real (Suc n))" by (intro continuous_intros) auto show "continuous_on {of_int a..of_int b} f" by fact show "(f has_vector_derivative f' t) (at t)" if "t \ {of_int a<..t. pbernpoly (Suc n) t *\<^sub>R f' t) integrable_on {a..b}" by (intro integrable_EM_remainder' cont') hence "(\t. (1 / real (Suc n)) *\<^sub>R pbernpoly (Suc n) t *\<^sub>R f' t) integrable_on {a..b}" by (rule integrable_cmul) also have "(\t. (1 / real (Suc n)) *\<^sub>R pbernpoly (Suc n) t *\<^sub>R f' t) = (\t. (pbernpoly (Suc n) t / real (Suc n)) *\<^sub>R f' t)" by (rule ext) (simp add: algebra_simps) finally show "((\t. (pbernpoly (Suc n) t / real (Suc n)) *\<^sub>R f' t) has_integral ?h b - ?h a - (- T + (?h b - ?h a))) {a..b}" using integrable_EM_remainder'[of a b f' "Suc n"] by (simp add: has_integral_iff T_def) qed (insert ab n, auto intro!: derivative_eq_intros simp: has_field_derivative_iff_has_vector_derivative [symmetric] not_le elim!: Ints_cases) also have "?h b - ?h a = (bernoulli (Suc n) / real (Suc n)) *\<^sub>R (f b - f a)" using n by (simp add: algebra_simps bernoulli'_def) finally have "integral {a..b} (\t. pbernpoly n t *\<^sub>R f t) = \ - T" by (simp add: has_integral_iff) also have "T = integral {a..b} (\t. (1 / real (Suc n)) *\<^sub>R (pbernpoly (Suc n) t) *\<^sub>R f' t)" by (simp add: T_def) also have "\ = integral {a..b} (\t. pbernpoly (Suc n) t *\<^sub>R f' t) /\<^sub>R real (Suc n)" by (subst integral_cmul) (simp_all add: divide_simps) finally show ?thesis . qed lemma EM_remainder'_conv_Suc: "EM_remainder' n f a b = ((-1) ^ Suc n * bernoulli (Suc n) / fact (Suc n)) *\<^sub>R (f b - f a) + EM_remainder' (Suc n) f' a b" by (simp add: EM_remainder'_def EM_remainder'_integral_conv_Suc scaleR_diff_right scaleR_add_right field_simps del: of_nat_Suc) end context fixes f f' :: "real \ 'a :: banach" and a :: int and n :: nat and A :: "real set" and C assumes n: "n > 0" assumes fin: "finite A" assumes cont: "continuous_on {of_int a..} f" assumes cont': "continuous_on {of_int a..} f'" assumes lim: "(f \ C) at_top" assumes deriv: "\x. x \ {of_int a<..} - A \ (f has_vector_derivative f' x) (at x)" begin lemma shows EM_remainder_converges_iff_Suc_converges: "EM_remainder_converges n f a \ EM_remainder_converges (Suc n) f' a" and EM_remainder_conv_Suc: "EM_remainder_converges n f a \ EM_remainder n f a = ((-1) ^ Suc n * bernoulli (Suc n) / fact (Suc n)) *\<^sub>R (C - f a) + EM_remainder (Suc n) f' a" proof (rule iffI) define g where "g = (\x. ((-1) ^ Suc n * bernoulli (Suc n) / fact (Suc n)) *\<^sub>R (f x - f a))" define G where "G = ((-1) ^ Suc n * bernoulli (Suc n) / fact (Suc n)) *\<^sub>R (C - f a)" have limit_g: "(g \ G) at_top" unfolding g_def G_def by (intro tendsto_intros lim) have *: "eventually (\x. EM_remainder' n f (real_of_int a) (real_of_int x) = g x + EM_remainder' (Suc n) f' (real_of_int a) (real_of_int x)) at_top" using eventually_ge_at_top[of a] proof eventually_elim case (elim b) thus ?case using EM_remainder'_conv_Suc[OF elim n fin continuous_on_subset[OF cont] continuous_on_subset[OF cont'] deriv] by (auto simp: g_def) qed { assume "EM_remainder_converges n f a" then obtain L where L: "((\b. EM_remainder' n f (real_of_int a) (real_of_int b)) \ L) at_top" by (auto simp: EM_remainder_converges_def) have *: "((\b. EM_remainder' (Suc n) f' (real_of_int a) (real_of_int b)) \ L - G) at_top" proof (rule Lim_transform_eventually) show "\\<^sub>F x in at_top. EM_remainder' n f (real_of_int a) (real_of_int x) - g x = EM_remainder' (Suc n) f' (real_of_int a) (real_of_int x)" using * by (simp add: algebra_simps) show "((\x. EM_remainder' n f (real_of_int a) (real_of_int x) - g x) \ L - G) at_top" by (intro tendsto_intros filterlim_compose[OF limit_g] L) qed from * show "EM_remainder_converges (Suc n) f' a" unfolding EM_remainder_converges_def .. from * have "EM_remainder (Suc n) f' a = L - G" by (rule EM_remainder_eqI) moreover from L have "EM_remainder n f a = L" by (rule EM_remainder_eqI) ultimately show "EM_remainder n f a = G + EM_remainder (Suc n) f' a" by (simp add: G_def) } { assume "EM_remainder_converges (Suc n) f' a" then obtain L where L: "((\b. EM_remainder' (Suc n) f' (real_of_int a) (real_of_int b)) \ L) at_top" by (auto simp: EM_remainder_converges_def) have *: "((\b. EM_remainder' n f (real_of_int a) (real_of_int b)) \ G + L) at_top" proof (rule Lim_transform_eventually) show "\\<^sub>F x in at_top. g x + EM_remainder' (Suc n) f' (real_of_int a) (real_of_int x) = EM_remainder' n f (real_of_int a) (real_of_int x)" using * by (subst eq_commute) show "((\x. g x + EM_remainder' (Suc n) f' (real_of_int a) (real_of_int x)) \ G + L) at_top" by (intro tendsto_intros filterlim_compose[OF limit_g] L) qed thus "EM_remainder_converges n f a" unfolding EM_remainder_converges_def .. } qed end subsection \The conventional version of the Euler--MacLaurin formula\ text \ The following theorems are the classic Euler--MacLaurin formula that can be found, with slight variations, in many sources (e.\,g.\ \cite{AS_HMF, apostol99, GKP_CM}). \ context fixes f :: "real \ 'a :: banach" fixes fs :: "nat \ real \ 'a" fixes a b :: int assumes ab: "a \ b" fixes N :: nat assumes N: "N > 0" fixes Y :: "real set" assumes fin: "finite Y" assumes fs_0 [simp]: "fs 0 = f" assumes fs_cont [continuous_intros]: "\k. k \ N \ continuous_on {real_of_int a..real_of_int b} (fs k)" assumes fs_deriv [derivative_intros]: "\k x. k < N \ x \ {a..b} - Y \ (fs k has_vector_derivative fs (Suc k) x) (at x)" begin theorem euler_maclaurin_raw_strong_int: defines "S \ (\i\{a<..b}. f (of_int i))" defines "I \ integral {of_int a..of_int b} f" defines "c' \ \k. (bernoulli' (Suc k) / fact (Suc k)) *\<^sub>R (fs k b - fs k a)" shows "S - I = (\k 'a" where "c = (\k. ((-1) ^ (Suc k) * bernoulli (Suc k) / fact (Suc k)) *\<^sub>R (fs k b - fs k a))" have "S - I = (\k 1" "m \ N" for m using that proof (induction m rule: dec_induct) case base with ab fin fs_cont[of 0] show ?case using fs_deriv[of 0] N unfolding One_nat_def by (subst EM_remainder'_Suc_0[of _ _ Y f]) (simp_all add: algebra_simps S_def I_def c_def) next case (step n) from step.prems have "S - I = (\kkkR (fs n b - fs n a))" (is "_ = _ + ?c") by (simp add: EM_remainder'_Suc_0 c_def) also have "\ + EM_remainder' n (fs n) a b = (\k b" "0 < N" "finite Y" "fs 0 = f" "(\k. k \ N \ continuous_on {real a..real b} (fs k))" "(\k x. k < N \ x \ {real a..real b} - Y \ (fs k has_vector_derivative fs (Suc k) x) (at x))" shows "(\i\{a<..b}. f (real i)) - integral {real a..real b} f = (\kR (fs k (real b) - fs k (real a))) + EM_remainder' N (fs N) (real a) (real b)" proof - have "(\i\{int a<..int b}. f (real_of_int i)) - integral {real_of_int (int a)..real_of_int (int b)} f = (\kR (fs k (real_of_int (int b)) - fs k (real_of_int (int a)))) + EM_remainder' N (fs N) (real_of_int (int a)) (real_of_int (int b))" using assms by (intro euler_maclaurin_raw_strong_int[where Y = Y] assms) simp_all also have "(\i\{int a<..int b}. f (real_of_int i)) = (\i\{a<..b}. f (real i))" by (intro sum.reindex_bij_witness[of _ int nat]) auto finally show ?thesis by simp qed subsection \The ``Concrete Mathematics'' version of the Euler--MacLaurin formula\ text \ As explained in \textit{Concrete Mathematics}~\cite{GKP_CM}, the above form of the formula has some drawbacks: When applying it to determine the asymptotics of some concrete function, one is usually left with several different unwieldy constant terms that are difficult to get rid of. There is no general way to determine what these constant terms are, but in concrete applications, they can often be determined or estimated by other means. We can therefore simply group all the constant terms into a single constant and have the user provide a proof of what it is. \ locale euler_maclaurin_int = fixes F f :: "real \ 'a :: banach" fixes fs :: "nat \ real \ 'a" fixes a :: int fixes N :: nat assumes N: "N > 0" fixes C :: 'a fixes Y :: "real set" assumes fin: "finite Y" assumes fs_0 [simp]: "fs 0 = f" assumes fs_cont [continuous_intros]: "\k. k \ N \ continuous_on {real_of_int a..} (fs k)" assumes fs_deriv [derivative_intros]: "\k x. k < N \ x \ {of_int a..} - Y \ (fs k has_vector_derivative fs (Suc k) x) (at x)" assumes F_cont [continuous_intros]: "continuous_on {of_int a..} F" assumes F_deriv [derivative_intros]: "\x. x \ {of_int a..} - Y \ (F has_vector_derivative f x) (at x)" assumes limit: "((\b. (\k=a..b. f k) - F (of_int b) - (\iR fs i (of_int b))) \ C) at_top" begin context fixes C' T defines "C' \ -f a + F a + C + (\kR (fs k (of_int a)))" and "T \ (\x. \iR fs i x)" begin lemma euler_maclaurin_strong_int_aux: assumes ab: "a \ b" defines "S \ (\k=a..b. f (of_int k))" shows "S - F (of_int b) - T (of_int b) = EM_remainder' N (fs N) (of_int a) (of_int b) + (C - C')" proof (cases "a = b") case True thus ?thesis unfolding C'_def by (simp add: S_def EM_remainder'_def T_def) next case False with assms have ab: "a < b" by simp define T' where "T' = (\kR (fs k (of_int a)))" have "(\i\{a<..b}. f (of_int i)) - integral {of_int a..of_int b} f = (\kR (fs k (of_int b) - fs k (of_int a))) + EM_remainder' N (fs N) (of_int a) (of_int b)" using ab by (intro euler_maclaurin_raw_strong_int [where Y = Y] N fin fs_0 continuous_on_subset[OF fs_cont] fs_deriv) auto also have "(f has_integral (F b - F a)) {of_int a..of_int b}" using ab by (intro fundamental_theorem_of_calculus_strong[OF fin]) (auto intro!: continuous_on_subset[OF F_cont] derivative_intros) hence "integral {of_int a..of_int b} f = F (of_int b) - F (of_int a)" by (simp add: has_integral_iff) also have "(\kR (fs k (of_int b) - fs k (of_int a))) = T (of_int b) - T'" by (simp add: T_def T'_def algebra_simps sum_subtractf) also have "(\i\{a<..b}. f (of_int i)) = S - f (of_int a)" unfolding S_def using ab by (subst sum_atLeastAtMost_int_head) auto finally show ?thesis by (simp add: algebra_simps C'_def T'_def) qed lemma EM_remainder_limit: assumes ab: "a \ b" defines "D \ EM_remainder' N (fs N) (of_int a) (of_int b)" shows "EM_remainder N (fs N) b = C' - D" and EM_remainder_converges: "EM_remainder_converges N (fs N) b" proof - note limit also have "((\b. (\k = a..b. f (of_int k)) - F (of_int b) - (\iR fs i (of_int b))) \ C) at_top = ((\b. (\k = a..b. f (of_int k)) - F (of_int b) - T (of_int b)) \ C) at_top" unfolding T_def .. also have "eventually (\x. (\k=a..x. f k) - F (of_int x) - T (of_int x) = EM_remainder' N (fs N) (of_int a) (of_int x) + (C - C')) at_top" (is "eventually (\x. ?f x = ?g x) _") using eventually_gt_at_top[of b] by eventually_elim (rule euler_maclaurin_strong_int_aux, insert ab, simp_all) hence "(?f \ C) at_top \ (?g \ C) at_top" by (intro filterlim_cong refl) finally have "((\x. ?g x - (C - C')) \ (C - (C - C'))) at_top" by (rule tendsto_diff[OF _ tendsto_const]) hence *: "((\x. EM_remainder' N (fs N) (of_int a) (of_int x)) \ C') at_top" by simp have "((\x. EM_remainder' N (fs N) (of_int a) (of_int x) - D) \ C' - D) at_top" by (intro tendsto_intros *) also have "eventually (\x. EM_remainder' N (fs N) (of_int a) (of_int x) - D = EM_remainder' N (fs N) (of_int b) (of_int x)) at_top" (is "eventually (\x. ?f x = ?g x) _") using eventually_ge_at_top[of b] proof eventually_elim case (elim x) have "EM_remainder' N (fs N) (of_int a) (of_int x) = D + EM_remainder' N (fs N) (of_int b) (of_int x)" using elim ab unfolding D_def by (intro EM_remainder'_combine [symmetric] continuous_on_subset[OF fs_cont]) auto thus ?case by simp qed hence "(?f \ C' - D) at_top \ (?g \ C' - D) at_top" by (intro filterlim_cong refl) finally have *: \ . from * show "EM_remainder_converges N (fs N) b" unfolding EM_remainder_converges_def .. from * show "EM_remainder N (fs N) b = C' - D" by (rule EM_remainder_eqI) qed theorem euler_maclaurin_strong_int: assumes ab: "a \ b" defines "S \ (\k=a..b. f (of_int k))" shows "S = F (of_int b) + C + T (of_int b) - EM_remainder N (fs N) b" proof - have "S = F (of_int b) + T (of_int b) + - (C' - EM_remainder' N (fs N) (of_int a) (of_int b)) + C" using euler_maclaurin_strong_int_aux[OF ab] by (simp add: algebra_simps S_def) also have "C' - EM_remainder' N (fs N) (of_int a) (of_int b) = EM_remainder N (fs N) b" using ab by (rule EM_remainder_limit(1) [symmetric]) finally show ?thesis by simp qed end end text \ The following version of the formula removes all the terms where the associated Bernoulli numbers vanish. \ locale euler_maclaurin_int' = fixes F f :: "real \ 'a :: banach" fixes fs :: "nat \ real \ 'a" fixes a :: int fixes N :: nat fixes C :: 'a fixes Y :: "real set" assumes fin: "finite Y" assumes fs_0 [simp]: "fs 0 = f" assumes fs_cont [continuous_intros]: "\k. k \ 2*N+1 \ continuous_on {real_of_int a..} (fs k)" assumes fs_deriv [derivative_intros]: "\k x. k \ 2*N \ x \ {of_int a..} - Y \ (fs k has_vector_derivative fs (Suc k) x) (at x)" assumes F_cont [continuous_intros]: "continuous_on {of_int a..} F" assumes F_deriv [derivative_intros]: "\x. x \ {of_int a..} - Y \ (F has_vector_derivative f x) (at x)" assumes limit: "((\b. (\k=a..b. f k) - F (of_int b) - (\i<2*N+1. (bernoulli' (Suc i) / fact (Suc i)) *\<^sub>R fs i (of_int b))) \ C) at_top" begin sublocale euler_maclaurin_int F f fs a "2*N+1" C Y by standard (insert fin fs_0 fs_cont fs_deriv F_cont F_deriv limit, simp_all) theorem euler_maclaurin_strong_int': assumes "a \ b" shows "(\k=a..b. f (of_int k)) = F (of_int b) + C + (1 / 2) *\<^sub>R f (of_int b) + (\i=1..N. (bernoulli (2*i) / fact (2*i)) *\<^sub>R fs (2*i-1) (of_int b)) - EM_remainder (2*N+1) (fs (2*N+1)) b" proof - have "(\k=a..b. f (real_of_int k)) = F (of_int b) + C + (\i<2*N+1. (bernoulli' (Suc i) / fact (Suc i)) *\<^sub>R fs i (of_int b)) - EM_remainder (2*N+1) (fs (2*N+1)) b" by (rule euler_maclaurin_strong_int) (simp_all only: lessThan_Suc_atMost Suc_eq_plus1 [symmetric] assms) also have "{..<2*N+1} = insert 0 {1..2*N}" by auto also have "(\i\\. (bernoulli' (Suc i) / fact (Suc i)) *\<^sub>R fs i (of_int b)) = (1 / 2) *\<^sub>R f (of_int b) + (\i\{1..2*N}. (bernoulli' (Suc i) / fact (Suc i)) *\<^sub>R fs i (of_int b))" by (subst sum.insert) (auto simp: assms bernoulli'_def) also have "(\i\{1..2*N}. (bernoulli' (Suc i) / fact (Suc i)) *\<^sub>R fs i (of_int b)) = (\i\{1..N}. (bernoulli' (2*i) / fact (2*i)) *\<^sub>R fs (2*i-1) (of_int b))" proof (rule sym, rule sum.reindex_bij_witness_not_neutral) fix i assume "i \ {1..2*N} - {i\{1..2*N}. even i}" thus "2 * ((i + 1) div 2) - 1 = i" "(i + 1) div 2 \ {1..N} - {}" by (auto elim!: oddE) qed (auto simp: bernoulli_odd_eq_0 bernoulli'_def algebra_simps) also have "\ = (\i\{1..N}. (bernoulli (2*i) / fact (2*i)) *\<^sub>R fs (2*i-1) (of_int b))" by (intro sum.cong refl) (auto simp: bernoulli'_def) finally show ?thesis by (simp only: add_ac) qed end text \ For convenience, we also offer a version where the sum ranges over natural numbers instead of integers. \ lemma sum_atLeastAtMost_of_int_nat_transfer: "(\k=int a..int b. f (of_int k)) = (\k=a..b. f (of_nat k))" by (intro sum.reindex_bij_witness[of _ int nat]) auto lemma euler_maclaurin_nat_int_transfer: fixes F and f :: "real \ 'a :: real_normed_vector" assumes "((\b. (\k=a..b. f (real k)) - F (real b) - T (real b)) \ C) at_top" shows "((\b. (\k=int a..b. f (of_int k)) - F (of_int b) - T (of_int b)) \ C) at_top" proof - have *: "((\b. (\k=int a..int b. f (of_int k)) - F (of_int (int b)) - T (of_int (int b))) \ C) at_top" using assms by (subst sum_atLeastAtMost_of_int_nat_transfer) simp thus ?thesis by (rule filterlim_int_of_nat_at_topD) qed locale euler_maclaurin_nat = fixes F f :: "real \ 'a :: banach" fixes fs :: "nat \ real \ 'a" fixes a :: nat fixes N :: nat assumes N: "N > 0" fixes C :: 'a fixes Y :: "real set" assumes fin: "finite Y" assumes fs_0 [simp]: "fs 0 = f" assumes fs_cont [continuous_intros]: "\k. k \ N \ continuous_on {real a..} (fs k)" assumes fs_deriv [derivative_intros]: "\k x. k < N \ x \ {real a..} - Y \ (fs k has_vector_derivative fs (Suc k) x) (at x)" assumes F_cont [continuous_intros]: "continuous_on {real a..} F" assumes F_deriv [derivative_intros]: "\x. x \ {real a..} - Y \ (F has_vector_derivative f x) (at x)" assumes limit: "((\b. (\k=a..b. f k) - F (real b) - (\iR fs i (real b))) \ C) at_top" begin sublocale euler_maclaurin_int F f fs "int a" N C Y by standard (insert N fin fs_cont fs_deriv F_cont F_deriv euler_maclaurin_nat_int_transfer[OF limit], simp_all) theorem euler_maclaurin_strong_nat: assumes ab: "a \ b" defines "S \ (\k=a..b. f (real k))" shows "S = F (real b) + C + (\iR fs i (real b)) - EM_remainder N (fs N) (int b)" using euler_maclaurin_strong_int[of "int b"] by (simp add: assms sum_atLeastAtMost_of_int_nat_transfer) end locale euler_maclaurin_nat' = fixes F f :: "real \ 'a :: banach" fixes fs :: "nat \ real \ 'a" fixes a :: nat fixes N :: nat fixes C :: 'a fixes Y :: "real set" assumes fin: "finite Y" assumes fs_0 [simp]: "fs 0 = f" assumes fs_cont [continuous_intros]: "\k. k \ 2*N+1 \ continuous_on {real a..} (fs k)" assumes fs_deriv [derivative_intros]: "\k x. k \ 2*N \ x \ {real a..} - Y \ (fs k has_vector_derivative fs (Suc k) x) (at x)" assumes F_cont [continuous_intros]: "continuous_on {real a..} F" assumes F_deriv [derivative_intros]: "\x. x \ {real a..} - Y \ (F has_vector_derivative f x) (at x)" assumes limit: "((\b. (\k=a..b. f k) - F (real b) - (\i<2*N+1. (bernoulli' (Suc i) / fact (Suc i)) *\<^sub>R fs i (real b))) \ C) at_top" begin sublocale euler_maclaurin_int' F f fs "int a" N C Y by standard (insert fin fs_cont fs_deriv F_cont F_deriv euler_maclaurin_nat_int_transfer[OF limit], simp_all) theorem euler_maclaurin_strong_nat': assumes "a \ b" shows "(\k=a..b. f (real k)) = F (real b) + C + (1 / 2) *\<^sub>R f (real b) + (\i=1..N. (bernoulli (2*i) / fact (2*i)) *\<^sub>R fs (2*i-1) (real b)) - EM_remainder (2*N+1) (fs (2*N+1)) b" using euler_maclaurin_strong_int'[of b] by (simp add: assms sum_atLeastAtMost_of_int_nat_transfer) end subsection \Bounds on the remainder term\ text \ The following theorems provide some simple means to bound the remainder terms. In practice, better bounds can often be obtained e.\,g. for the $n$-th remainder term by expanding it to the sum of the first discarded term in the expansion and the $n+1$-th remainder term. \ lemma fixes f :: "real \ 'a :: {real_normed_field, banach}" and g g' :: "real \ real" assumes fin: "finite Y" assumes pbernpoly_bound: "\x. \pbernpoly n x\ \ D" assumes cont_f: "continuous_on {a..} f" assumes cont_g: "continuous_on {a..} g" assumes cont_g': "continuous_on {a..} g'" assumes limit_g: "(g \ C) at_top" assumes f_bound: "\x. x \ a \ norm (f x) \ g' x" assumes deriv: "\x. x \ {a..} - Y \ (g has_field_derivative g' x) (at x)" shows norm_EM_remainder_le_strong_int: "\x. of_int x \ a \ norm (EM_remainder n f x) \ D / fact n * (C - g x)" and norm_EM_remainder_le_strong_nat: "\x. real x \ a \ norm (EM_remainder n f (int x)) \ D / fact n * (C - g x)" proof - from pbernpoly_bound have D: "norm (pbernpoly n x) \ D" for x by auto from this[of 0] have D_nonneg: "D \ 0" by simp define D' where "D' = D / fact n" from D_nonneg have D'_nonneg: "D' \ 0" by (simp add: D'_def) have bound: "norm (EM_remainder' n f x y) \ D' * (g y - g x)" if xy: "x \ a" "x \ y" for x y :: real proof - have "norm (EM_remainder' n f x y) = norm (integral {x..y} (\t. pbernpoly n t *\<^sub>R f t)) / fact n" by (simp add: EM_remainder'_def) also have "(\t. D * g' t) integrable_on {x..y}" using xy by (intro integrable_continuous_real continuous_intros continuous_on_subset[OF cont_g']) auto hence "norm (integral {x..y} (\t. pbernpoly n t *\<^sub>R f t)) \ integral {x..y} (\t. D * g' t)" using D D_nonneg xy by (intro integral_norm_bound_integral integrable_EM_remainder' continuous_on_subset[OF cont_f]) (auto intro!: mult_mono f_bound) also have "\ = D * integral {x..y} g'" by simp also have "(g' has_integral (g y - g x)) {x..y}" using xy by (intro fundamental_theorem_of_calculus_strong[OF fin] continuous_on_subset[OF cont_g]) (auto simp: has_field_derivative_iff_has_vector_derivative [symmetric] intro!: deriv) hence "integral {x..y} g' = g y - g x" by (simp add: has_integral_iff) finally show ?thesis by (simp add: D'_def divide_simps) qed have lim: "((\y. EM_remainder' n f x (of_int y)) \ EM_remainder n f x) at_top" if x: "x \ a" for x :: int proof - have "(\n. g (real n)) \ C" by (rule filterlim_compose[OF limit_g filterlim_real_sequentially]) hence Cauchy: "Cauchy (\n. g (real n))" using convergent_eq_Cauchy by blast have "Cauchy (\y. EM_remainder' n f x (int y))" proof (rule CauchyI', goal_cases) case (1 \) define \' where "\' = (if D' = 0 then 1 else \ / (2*D'))" from \\ > 0\ D'_nonneg have \': "\' > 0" by (simp add: \'_def divide_simps) from CauchyD[OF Cauchy this] obtain M where M: "\m n. m \ M \ n \ M \ norm (g (real m) - g (real n)) < \'" by blast show ?case proof (intro CauchyI' exI[of _ "max (max 0 M) (nat x)"] allI impI, goal_cases) case (1 k l) have "EM_remainder' n f x k + EM_remainder' n f k l = EM_remainder' n f x l" using 1 x by (intro EM_remainder'_combine continuous_on_subset[OF cont_f]) auto hence "EM_remainder' n f x l - EM_remainder' n f x k = EM_remainder' n f k l" by (simp add: algebra_simps) also from 1 x have "norm \ \ D' * (g l - g k)" by (intro bound) auto also have "g l - g k \ norm (g l - g k)" by simp also from 1 have "\ \ \'" using M[of l k] by auto also from \\ > 0\ have "D' * \' \ \ / 2" by (simp add: \'_def) also from \\ > 0\ have "\ < \" by simp finally show ?case by (simp add: D'_nonneg mult_left_mono dist_norm norm_minus_commute) qed qed then obtain L where "(\y. EM_remainder' n f x (int y)) \ L" by (auto simp: Cauchy_convergent_iff convergent_def) from filterlim_int_of_nat_at_topD[OF this] have "((\y. EM_remainder' n f x (of_int y)) \ L) at_top" by simp moreover from this have "EM_remainder n f x = L" by (rule EM_remainder_eqI) ultimately show "((\y. EM_remainder' n f x (of_int y)) \ EM_remainder n f x) at_top" by simp qed have *: "norm (EM_remainder n f x) \ D' * (C - g x)" if x: "x \ a" for x :: int proof (rule tendsto_le) show "((\y. D' * (g (of_int y) - g (of_int x))) \ D' * (C - g (of_int x))) at_top" by (intro tendsto_intros filterlim_compose[OF limit_g]) show "((\y. norm (EM_remainder' n f x (of_int y))) \ norm (EM_remainder n f x)) at_top" using x by (intro tendsto_norm lim) show "eventually (\y. norm (EM_remainder' n f (of_int x) (of_int y)) \ D' * (g (of_int y) - g (of_int x))) at_top" using eventually_ge_at_top[of x] by eventually_elim (rule bound, insert x, simp_all) qed simp_all thus "\x. of_int x \ a \ norm (EM_remainder n f x) \ D' * (C - g x)" by blast have "norm (EM_remainder n f x) \ D' * (C - g x)" if x: "x \ a" for x :: nat using x *[of "int x"] by simp thus "\x. real x \ a \ norm (EM_remainder n f (int x)) \ D' * (C - g x)" by blast qed lemma fixes f :: "real \ 'a :: {real_normed_field, banach}" and g g' :: "real \ real" assumes fin: "finite Y" assumes pbernpoly_bound: "\x. \pbernpoly n x\ \ D" assumes cont_f: "continuous_on {a..} f" assumes cont_g: "continuous_on {a..} g" assumes cont_g': "continuous_on {a..} g'" assumes limit_g: "(g \ 0) at_top" assumes f_bound: "\x. x \ a \ norm (f x) \ g' x" assumes deriv: "\x. x \ {a..} - Y \ (g has_field_derivative -g' x) (at x)" shows norm_EM_remainder_le_strong_int': "\x. of_int x \ a \ norm (EM_remainder n f x) \ D / fact n * g x" and norm_EM_remainder_le_strong_nat': "\x. real x \ a \ norm (EM_remainder n f (int x)) \ D / fact n * g x" proof - have "\x. of_int x \ a \ norm (EM_remainder n f x) \ D / fact n * (0 - (-g x))" using assms by (intro norm_EM_remainder_le_strong_int[OF fin pbernpoly_bound _ _ cont_g']) (auto intro!: continuous_intros tendsto_eq_intros derivative_eq_intros) thus "\x. of_int x \ a \ norm (EM_remainder n f x) \ D / fact n * g x" by auto next have "\x. real x \ a \ norm (EM_remainder n f (int x)) \ D / fact n * (0 - (-g x))" using assms by (intro norm_EM_remainder_le_strong_nat[OF fin pbernpoly_bound _ _ cont_g']) (auto intro!: continuous_intros tendsto_eq_intros derivative_eq_intros) thus "\x. real x \ a \ norm (EM_remainder n f (int x)) \ D / fact n * g x" by auto qed lemma norm_EM_remainder'_le: fixes f :: "real \ 'a :: {real_normed_field, banach}" and g g' :: "real \ real" assumes cont_f: "continuous_on {a..} f" assumes cont_g': "continuous_on {a..} g'" assumes f_bigo: "eventually (\x. norm (f x) \ g' x) at_top" assumes deriv: "eventually (\x. (g has_field_derivative g' x) (at x)) at_top" obtains C D where "eventually (\x. norm (EM_remainder' n f a x) \ C + D * g x) at_top" proof - note cont = continuous_on_subset[OF cont_f] continuous_on_subset[OF cont_g'] from bounded_pbernpoly[of n] obtain D where D: "\x. norm (pbernpoly n x) \ D" by blast from this[of 0] have D_nonneg: "D \ 0" by simp from eventually_conj[OF f_bigo eventually_conj[OF deriv eventually_ge_at_top[of a]]] obtain x0 where x0: "x0 \ a" "\x. x \ x0 \ norm (f x) \ g' x" "\x. x \ x0 \ (g has_field_derivative g' x) (at x)" by (auto simp: eventually_at_top_linorder) define C where "C = (norm (integral {a..x0} (\t. pbernpoly n t *\<^sub>R f t)) - D * g x0) / fact n" have "eventually (\x. norm (EM_remainder' n f a x) \ C + D / fact n * g x) at_top" using eventually_ge_at_top[of x0] proof eventually_elim case (elim x) have "integral {a..x} (\t. pbernpoly n t *\<^sub>R f t) = integral {a..x0} (\t. pbernpoly n t *\<^sub>R f t) + integral {x0..x} (\t. pbernpoly n t *\<^sub>R f t)" (is "_ = ?I1 + ?I2") using elim x0(1) by (intro Henstock_Kurzweil_Integration.integral_combine [symmetric] integrable_EM_remainder' cont) auto also have "norm \ \ norm ?I1 + norm ?I2" by (rule norm_triangle_ineq) also have "norm ?I2 \ integral {x0..x} (\t. D * g' t)" using x0 D D_nonneg by (intro integral_norm_bound_integral integrable_EM_remainder') (auto intro!: integrable_continuous_real continuous_intros cont mult_mono) also have "\ = D * integral {x0..x} g'" by simp also from elim have "(g' has_integral (g x - g x0)) {x0..x}" by (intro fundamental_theorem_of_calculus) (auto intro!: has_field_derivative_at_within[OF x0(3)] simp: has_field_derivative_iff_has_vector_derivative [symmetric]) hence "integral {x0..x} g' = g x - g x0" by (simp add: has_integral_iff) finally have "norm (integral {a..x} (\t. pbernpoly n t *\<^sub>R f t)) \ norm ?I1 + D * (g x - g x0)" by simp_all thus "norm (EM_remainder' n f a x) \ C + D / fact n * g x" by (simp add: EM_remainder'_def field_simps C_def) qed thus ?thesis by (rule that) qed subsection \Application to harmonic numbers\ text \ As a first application, we can apply the machinery we have developed to the harmonic numbers. \ definition harm_remainder :: "nat \ nat \ real" where "harm_remainder N n = EM_remainder (2*N+1) (\x. -fact (2*N+1) / x ^ (2*N+2)) (int n)" lemma harm_expansion: assumes n: "n > 0" and N: "N > 0" shows "harm n = ln n + euler_mascheroni + 1 / (2*n) - (\i=1..N. bernoulli (2*i) / ((2*i) * n ^ (2*i))) - harm_remainder N n" proof - define fs where "fs = (\k x. (-1) ^ k * fact k / x ^ (Suc k) :: real)" interpret euler_maclaurin_nat' ln "\x. 1/x" fs 1 N euler_mascheroni "{}" proof fix k x assume "k \ 2*N" "x \ {real 1..} - {}" thus "(fs k has_vector_derivative fs (Suc k) x) (at x)" by (cases "k = 0") (auto intro!: derivative_eq_intros simp: fs_def has_field_derivative_iff_has_vector_derivative [symmetric] field_simps power_diff) next have "(\b. harm b - ln (real b) - (\i<2*N+1. bernoulli' (Suc i) * (- 1) ^ i / (real (Suc i) * (real b ^ Suc i)))) \ (euler_mascheroni - (\i<2*N+1. 0))" by (intro tendsto_diff euler_mascheroni_LIMSEQ tendsto_sum real_tendsto_divide_at_top[OF tendsto_const] filterlim_tendsto_pos_mult_at_top[OF tendsto_const] filterlim_pow_at_top filterlim_real_sequentially) auto thus "(\b. (\k = 1..b. 1 / real k) - ln (real b) - (\i<2*N+1. (bernoulli' (Suc i) / fact (Suc i)) *\<^sub>R fs i (real b))) \ euler_mascheroni" by (simp add: harm_def divide_simps fs_def) qed (insert n N, auto intro!: continuous_intros derivative_eq_intros simp: fs_def has_field_derivative_iff_has_vector_derivative [symmetric]) have "harm n = (\k=1..n. 1 / real k)" by (simp add: harm_def divide_simps) also have "\ = ln (real n) + euler_mascheroni + (1/2) *\<^sub>R (1 / real n) + (\i=1..N. (bernoulli (2*i) / fact (2*i)) *\<^sub>R fs (2*i-1) (real n)) - EM_remainder (2*N+1) (fs (2*N+1)) (int n)" using n N using n by (intro euler_maclaurin_strong_nat') simp_all also have "(\i=1..N. (bernoulli (2*i) / fact (2*i)) *\<^sub>R (fs (2*i-1) (real n))) = (\i=1..N. -(bernoulli (2*i) / (real (2*i) * real n ^ (2*i))))" by (intro sum.cong refl) (simp_all add: fs_def divide_simps fact_reduce del: of_nat_Suc power_Suc) also have "\ = -(\i=1..N. bernoulli (2*i) / (real (2*i) * real n ^ (2*i)))" by (simp add: sum_negf) finally show ?thesis unfolding fs_def by (simp add: harm_remainder_def) qed lemma of_nat_ge_1_iff: "of_nat x \ (1 :: 'a :: linordered_semidom) \ x \ 1" using of_nat_le_iff[of 1 x] by (simp del: of_nat_le_iff) lemma harm_remainder_bound: fixes N :: nat assumes N: "N > 0" shows "\C. \n\1. norm (harm_remainder N n) \ C / real n ^ (2*N+1)" proof - from bounded_pbernpoly[of "2*N+1"] obtain D where D: "\x. \pbernpoly (2*N+1) x\ \ D" by auto have "\x. 1 \ real x \ norm (harm_remainder N x) \ D / fact (2*N+1) * (fact (2*N) / x ^ (2*N+1))" unfolding harm_remainder_def of_int_of_nat_eq proof (rule norm_EM_remainder_le_strong_nat'[of "{}"]) fix x :: real assume x: "x \ 1" show "norm (-fact (2*N+1) / x ^ (2 * N + 2)) \ fact (2*N+1) / x ^ (2*N+2)" using x by simp next show "((\x::real. fact (2 * N) / x ^ (2 * N + 1)) \ 0) at_top" by (intro real_tendsto_divide_at_top[OF tendsto_const] filterlim_pow_at_top filterlim_ident) simp_all qed (insert N D, auto intro!: derivative_eq_intros continuous_intros simp: field_simps power_diff) hence "\x. 1 \ x \ norm (harm_remainder N x) \ D / (2*N+1) / real x ^ (2*N+1)" by simp thus ?thesis by blast qed subsection \Application to sums of inverse squares\ text \ In the same vein, we can derive the asymptotics of the partial sum of inverse squares. \ lemma sum_inverse_squares_expansion: assumes n: "n > 0" and N: "N > 0" shows "(\k=1..n. 1 / real k ^ 2) = pi ^ 2 / 6 - 1 / real n + 1 / (2 * real n ^ 2) - (\i=1..N. bernoulli (2*i) / n ^ (2*i+1)) - EM_remainder (2*N+1) (\x. -fact (2*N+2) / x ^ (2*N+3)) (int n)" proof - have 3: "3 = Suc (Suc (Suc 0))" by (simp add: eval_nat_numeral) define fs where "fs = (\k x. (-1) ^ k * fact (Suc k) / x ^ (k+2) :: real)" interpret euler_maclaurin_nat' "\x. -1/x" "\x. 1/x^2" fs 1 N "pi^2/6" "{}" proof fix k x assume "k \ 2*N" "x \ {real 1..} - {}" thus "(fs k has_vector_derivative fs (Suc k) x) (at x)" by (cases "k = 0") (auto intro!: derivative_eq_intros simp: fs_def has_field_derivative_iff_has_vector_derivative [symmetric] field_simps power_diff) next from inverse_squares_sums have "(\n. \k pi\<^sup>2 / 6" by (simp add: sums_def) also have "(\n. \kn. \k=1..n. 1 / real k ^ 2)" by (intro ext sum.reindex_bij_witness[of _ "\n. n - 1" Suc]) auto finally have "(\b. (\k = 1..b. 1 / real k^2) + 1 / real b - (\i<2*N+1. bernoulli' (Suc i) * (- 1) ^ i / (real b ^ (i+2)))) \ (pi^2/6 + 0 - (\i<2*N+1. 0))" by (intro tendsto_diff tendsto_add real_tendsto_divide_at_top[OF tendsto_const] filterlim_tendsto_pos_mult_at_top[OF tendsto_const] filterlim_pow_at_top filterlim_real_sequentially tendsto_sum) auto thus "(\b. (\k = 1..b. 1 / real k^2) - (- 1 / real b) - (\i<2*N+1. (bernoulli' (Suc i) / fact (Suc i)) *\<^sub>R fs i (real b))) \ pi^2/6" by (simp add: harm_def field_simps fs_def del: power_Suc of_nat_Suc) qed (insert n N, auto intro!: continuous_intros derivative_eq_intros simp: fs_def has_field_derivative_iff_has_vector_derivative [symmetric] power2_eq_square) have "(\k=1..n. 1 / real k ^ 2) = - 1 / real n + pi^2/6 + (1/2) *\<^sub>R (1 / real n^2) + (\i=1..N. (bernoulli (2*i) / fact (2*i)) *\<^sub>R fs (2*i-1) (real n)) - EM_remainder (2*N+1) (fs (2*N+1)) (int n)" using n N using n by (intro euler_maclaurin_strong_nat') simp_all also have "(\i=1..N. (bernoulli (2*i) / fact (2*i)) *\<^sub>R (fs (2*i-1) (real n))) = (\i=1..N. -(bernoulli (2*i) / (real n ^ (2*i+1))))" by (intro sum.cong refl) (simp_all add: fs_def divide_simps fact_reduce del: of_nat_Suc power_Suc) also have "\ = -(\i=1..N. bernoulli (2*i) / real n ^ (2*i+1))" by (simp add: sum_negf) finally show ?thesis unfolding fs_def by (simp add: fs_def 3) qed lemma sum_inverse_squares_remainder_bound: fixes N :: nat assumes N: "N > 0" defines "R \ (\n. EM_remainder (2*N+1) (\x. -fact (2*N+2) / x ^ (2*N+3)) (int n))" shows "\C. \n\1. norm (R n) \ C / real n ^ (2*N+2)" proof - have 3: "3 = Suc (Suc (Suc 0))" by simp from bounded_pbernpoly[of "2*N+1"] obtain D where D: "\x. \pbernpoly (2*N+1) x\ \ D" by auto have "\x. 1 \ real x \ norm (R x) \ D / fact (2*N+1) * (fact (2*N+1) / x ^ (2*N+2))" unfolding R_def of_int_of_nat_eq proof (rule norm_EM_remainder_le_strong_nat'[of "{}"]) fix x :: real assume x: "x \ 1" show "norm (-fact (2*N+2) / x ^ (2*N+3)) \ fact (2*N+2) / x ^ (2*N+3)" using x by simp next show "((\x::real. fact (2*N+1) / x ^ (2*N+2)) \ 0) at_top" by (intro real_tendsto_divide_at_top[OF tendsto_const] filterlim_pow_at_top filterlim_ident) simp_all qed (insert N D, auto intro!: derivative_eq_intros continuous_intros simp: field_simps power_diff 3) hence "\x\1. norm (R x) \ D / real x ^ (2 * N + 2)" by simp thus ?thesis by blast qed end diff --git a/thys/Euler_MacLaurin/ROOT b/thys/Euler_MacLaurin/ROOT --- a/thys/Euler_MacLaurin/ROOT +++ b/thys/Euler_MacLaurin/ROOT @@ -1,13 +1,13 @@ chapter AFP -session Euler_MacLaurin (AFP) = "HOL-Analysis" + +session Euler_MacLaurin (AFP) = Bernoulli + options [timeout = 600] sessions - Bernoulli + "HOL-Complex_Analysis" Landau_Symbols theories Euler_MacLaurin Euler_MacLaurin_Landau document_files "root.tex" "root.bib" diff --git a/thys/Green/Derivs.thy b/thys/Green/Derivs.thy --- a/thys/Green/Derivs.thy +++ b/thys/Green/Derivs.thy @@ -1,654 +1,651 @@ theory Derivs imports General_Utils begin lemma field_simp_has_vector_derivative [derivative_intros]: "(f has_field_derivative y) F \ (f has_vector_derivative y) F" by (simp add: has_field_derivative_iff_has_vector_derivative) lemma continuous_on_cases_empty [continuous_intros]: "\closed S; continuous_on S f; \x. \x \ S; \ P x\ \ f x = g x\ \ continuous_on S (\x. if P x then f x else g x)" using continuous_on_cases [of _ "{}"] by force lemma inj_on_cases: assumes "inj_on f (Collect P \ S)" "inj_on g (Collect (Not \ P) \ S)" "f ` (Collect P \ S) \ g ` (Collect (Not \ P) \ S) = {}" shows "inj_on (\x. if P x then f x else g x) S" using assms by (force simp: inj_on_def) lemma inj_on_arccos: "S \ {-1..1} \ inj_on arccos S" by (metis atLeastAtMost_iff cos_arccos inj_onI subsetCE) lemma has_vector_derivative_componentwise_within: "(f has_vector_derivative f') (at a within S) \ (\i \ Basis. ((\x. f x \ i) has_vector_derivative (f' \ i)) (at a within S))" apply (simp add: has_vector_derivative_def) apply (subst has_derivative_componentwise_within) apply simp done lemma has_vector_derivative_pair_within: fixes f :: "real \ 'a::euclidean_space" and g :: "real \ 'b::euclidean_space" assumes "\u. u \ Basis \ ((\x. f x \ u) has_vector_derivative f' \ u) (at x within S)" "\u. u \ Basis \ ((\x. g x \ u) has_vector_derivative g' \ u) (at x within S)" shows "((\x. (f x, g x)) has_vector_derivative (f',g')) (at x within S)" apply (subst has_vector_derivative_componentwise_within) apply (auto simp: assms Basis_prod_def) done lemma piecewise_C1_differentiable_const: shows "(\x. c) piecewise_C1_differentiable_on s" using continuous_on_const by (auto simp add: piecewise_C1_differentiable_on_def) declare piecewise_C1_differentiable_const [simp, derivative_intros] declare piecewise_C1_differentiable_neg [simp, derivative_intros] declare piecewise_C1_differentiable_add [simp, derivative_intros] declare piecewise_C1_differentiable_diff [simp, derivative_intros] -(*move to Cauchy_Integral_Theorem*) +(*move to Derivative*) lemma piecewise_C1_differentiable_on_ident [simp, derivative_intros]: fixes f :: "real \ 'a::real_normed_vector" shows "(\x. x) piecewise_C1_differentiable_on S" unfolding piecewise_C1_differentiable_on_def using C1_differentiable_on_ident by (blast intro: continuous_on_id C1_differentiable_on_ident) lemma piecewise_C1_differentiable_on_mult [simp, derivative_intros]: fixes f :: "real \ 'a::real_normed_algebra" assumes "f piecewise_C1_differentiable_on S" "g piecewise_C1_differentiable_on S" shows "(\x. f x * g x) piecewise_C1_differentiable_on S" using assms unfolding piecewise_C1_differentiable_on_def apply safe apply (blast intro: continuous_intros) apply (rename_tac A B) apply (rule_tac x="A \ B" in exI) apply (auto intro: C1_differentiable_on_mult C1_differentiable_on_subset) done lemma C1_differentiable_on_cdiv [simp, derivative_intros]: fixes f :: "real \ 'a :: real_normed_field" shows "f C1_differentiable_on S \ (\x. f x / c) C1_differentiable_on S" by (simp add: divide_inverse) lemma piecewise_C1_differentiable_on_cdiv [simp, derivative_intros]: fixes f :: "real \ 'a::real_normed_field" assumes "f piecewise_C1_differentiable_on S" shows "(\x. f x / c) piecewise_C1_differentiable_on S" by (simp add: divide_inverse piecewise_C1_differentiable_const piecewise_C1_differentiable_on_mult assms) lemma sqrt_C1_differentiable [simp, derivative_intros]: assumes f: "f C1_differentiable_on S" and fim: "f ` S \ {0<..}" shows "(\x. sqrt (f x)) C1_differentiable_on S" proof - have contf: "continuous_on S f" by (simp add: C1_differentiable_imp_continuous_on f) show ?thesis using assms unfolding C1_differentiable_on_def has_field_derivative_iff_has_vector_derivative [symmetric] by (fastforce intro!: contf continuous_intros derivative_intros) qed lemma sqrt_piecewise_C1_differentiable [simp, derivative_intros]: assumes f: "f piecewise_C1_differentiable_on S" and fim: "f ` S \ {0<..}" shows "(\x. sqrt (f x)) piecewise_C1_differentiable_on S" using assms unfolding piecewise_C1_differentiable_on_def by (fastforce intro!: continuous_intros derivative_intros) lemma fixes f :: "real \ 'a::{banach,real_normed_field}" assumes f: "f C1_differentiable_on S" shows sin_C1_differentiable [simp, derivative_intros]: "(\x. sin (f x)) C1_differentiable_on S" and cos_C1_differentiable [simp, derivative_intros]: "(\x. cos (f x)) C1_differentiable_on S" proof - have contf: "continuous_on S f" by (simp add: C1_differentiable_imp_continuous_on f) note df_sin = field_vector_diff_chain_at [where g=sin, unfolded o_def] note df_cos = field_vector_diff_chain_at [where g=cos, unfolded o_def] show "(\x. sin (f x)) C1_differentiable_on S" "(\x. cos (f x)) C1_differentiable_on S" using assms unfolding C1_differentiable_on_def has_field_derivative_iff_has_vector_derivative [symmetric] apply auto by (rule contf continuous_intros df_sin df_cos derivative_intros exI conjI ballI | force)+ qed lemma has_derivative_abs: fixes a::real assumes "a \ 0" shows "(abs has_derivative ((*) (sgn a))) (at a)" proof - have [simp]: "norm = abs" using real_norm_def by force show ?thesis using has_derivative_norm [where 'a=real, simplified] assms by (simp add: mult_commute_abs) qed lemma abs_C1_differentiable [simp, derivative_intros]: fixes f :: "real \ real" assumes f: "f C1_differentiable_on S" and "0 \ f ` S" shows "(\x. abs (f x)) C1_differentiable_on S" proof - have contf: "continuous_on S f" by (simp add: C1_differentiable_imp_continuous_on f) note df = DERIV_chain [where f=abs and g=f, unfolded o_def] show ?thesis using assms unfolding C1_differentiable_on_def has_field_derivative_iff_has_vector_derivative [symmetric] apply clarify apply (rule df exI conjI ballI)+ apply (force simp: has_field_derivative_def intro: has_derivative_abs continuous_intros contf)+ done qed lemma C1_differentiable_on_pair [simp, derivative_intros]: fixes f :: "real \ 'a::euclidean_space" and g :: "real \ 'b::euclidean_space" assumes "f C1_differentiable_on S" "g C1_differentiable_on S" shows "(\x. (f x, g x)) C1_differentiable_on S" using assms unfolding C1_differentiable_on_def apply safe apply (rename_tac A B) apply (intro exI ballI conjI) apply (rule_tac f'="A x" and g'="B x" in has_vector_derivative_pair_within) using has_vector_derivative_componentwise_within by (blast intro: continuous_on_Pair)+ lemma piecewise_C1_differentiable_on_pair [simp, derivative_intros]: fixes f :: "real \ 'a::euclidean_space" and g :: "real \ 'b::euclidean_space" assumes "f piecewise_C1_differentiable_on S" "g piecewise_C1_differentiable_on S" shows "(\x. (f x, g x)) piecewise_C1_differentiable_on S" using assms unfolding piecewise_C1_differentiable_on_def by (blast intro!: continuous_intros C1_differentiable_on_pair del: continuous_on_discrete intro: C1_differentiable_on_subset) lemma test2: assumes s: "\x. x \ {0..1} - s \ g differentiable at x" and fs: "finite s" and uv: "u \ {0..1}" "v \ {0..1}" "u \ v" and "x \ {0..1}" "x \ (\t. (v-u) *\<^sub>R t + u) -` s" shows "vector_derivative (\x. g ((v-u) * x + u)) (at x within {0..1}) = (v-u) *\<^sub>R vector_derivative g (at ((v-u) * x + u) within{0..1})" proof - have i:"(g has_vector_derivative vector_derivative g (at ((v - u) * x + u))) (at ((v-u) * x + u))" using assms s [of "(v - u) * x + u"] uv mult_left_le [of x "v-u"] by (auto simp: vector_derivative_works) have ii:"((\x. g ((v - u) * x + u)) has_vector_derivative (v - u) *\<^sub>R vector_derivative g (at ((v - u) * x + u))) (at x)" by (intro vector_diff_chain_at [simplified o_def] derivative_eq_intros | simp add: i)+ have 0: "0 \ (v - u) * x + u" using assms uv by auto have 1: "(v - u) * x + u \ 1" using assms uv by simp (metis add.commute atLeastAtMost_iff atLeastatMost_empty_iff diff_ge_0_iff_ge empty_iff le_diff_eq mult_left_le) have iii: "vector_derivative g (at ((v - u) * x + u) within {0..1}) = vector_derivative g (at ((v - u) * x + u))" using Derivative.vector_derivative_at_within_ivl[OF i, of "0" "1", OF 0 1] by auto have iv: "vector_derivative (\x. g ((v - u) * x + u)) (at x within {0..1}) = (v - u) *\<^sub>R vector_derivative g (at ((v - u) * x + u))" using Derivative.vector_derivative_at_within_ivl[OF ii, of "0" "1"] assms by auto show ?thesis using iii iv by auto qed lemma C1_differentiable_on_components: assumes "\i. i \ Basis \ (\x. f x \ i) C1_differentiable_on s" shows "f C1_differentiable_on s" proof (clarsimp simp add: C1_differentiable_on_def has_vector_derivative_def) have *:"\f i x. x *\<^sub>R (f \ i) = (x *\<^sub>R f) \ i" by auto have "\f'. \i\Basis. \x\s. ((\x. f x \ i) has_derivative (\z. z *\<^sub>R f' x \ i)) (at x) \ continuous_on s f'" using assms lambda_skolem_euclidean[of "\i D. (\x\s. ((\x. f x \ i) has_derivative (\z. z *\<^sub>R D x)) (at x)) \ continuous_on s D"] apply (simp only: C1_differentiable_on_def has_vector_derivative_def *) using continuous_on_componentwise[of "s"] by metis then obtain f' where f': "\i\Basis. \x\s. ((\x. f x \ i) has_derivative (\z. z *\<^sub>R f' x \ i)) (at x) \ continuous_on s f'" by auto then have 0: "(\x\s. (f has_derivative (\z. z *\<^sub>R f' x)) (at x)) \ continuous_on s f'" using f' has_derivative_componentwise_within[of "f", where S= UNIV] by auto then show "\D. (\x\s. (f has_derivative (\z. z *\<^sub>R D x)) (at x)) \ continuous_on s D" by metis qed lemma piecewise_C1_differentiable_on_components: assumes "finite t" "\i. i \ Basis \ (\x. f x \ i) C1_differentiable_on s - t" "\i. i \ Basis \ continuous_on s (\x. f x \ i)" shows "f piecewise_C1_differentiable_on s" using C1_differentiable_on_components assms continuous_on_componentwise piecewise_C1_differentiable_on_def by blast lemma all_components_smooth_one_pw_smooth_is_pw_smooth: assumes "\i. i \ Basis - {j} \ (\x. f x \ i) C1_differentiable_on s" assumes "(\x. f x \ j) piecewise_C1_differentiable_on s" shows "f piecewise_C1_differentiable_on s" proof - have is_cont: "\i\Basis. continuous_on s (\x. f x \ i)" using assms C1_differentiable_imp_continuous_on piecewise_C1_differentiable_on_def by fastforce obtain t where t:"(finite t \ (\x. f x \ j) C1_differentiable_on s - t)" using assms(2) piecewise_C1_differentiable_on_def by auto show ?thesis using piecewise_C1_differentiable_on_components[where ?f = "f"] using assms(2) piecewise_C1_differentiable_on_def C1_differentiable_on_subset[OF assms(1) Diff_subset, where ?B1 ="t"] t is_cont by fastforce qed lemma derivative_component_fun_component: fixes i::"'a::euclidean_space" assumes "f differentiable (at x)" shows "((vector_derivative f (at x)) \ i) = ((vector_derivative (\x. (f x) \ i) (at x)) )" proof - have "((\x. f x \ i) has_vector_derivative vector_derivative f (at x) \ i) (at x)" using assms and bounded_linear.has_vector_derivative[of "(\x. x \ i)" "f" "(vector_derivative f (at x))" "(at x)"] and bounded_linear_inner_left[of "i"] and vector_derivative_works[of "f" "(at x)"] by blast then show "((vector_derivative f (at x)) \ i) = ((vector_derivative (\x. (f x) \ i) (at x)) )" using vector_derivative_works[of "(\x. (f x) \ i)" "(at x)"] and differentiableI_vector[of "(\x. (f x) \ i)" "(vector_derivative f (at x) \ i)" "(at x)"] and Derivative.vector_derivative_at by force qed lemma gamma_deriv_at_within: assumes a_leq_b: "a < b" and x_within_bounds: "x \ {a..b}" and gamma_differentiable: "\x \ {a .. b}. \ differentiable at x" shows "vector_derivative \ (at x within {a..b}) = vector_derivative \ (at x)" using Derivative.vector_derivative_at_within_ivl[of "\" "vector_derivative \ (at x)" "x" "a" "b"] gamma_differentiable x_within_bounds a_leq_b by (auto simp add: vector_derivative_works) lemma islimpt_diff_finite: assumes "finite (t::'a::t1_space set)" shows "x islimpt s - t = x islimpt s" proof- have iii: "s - t = s - (t \ s)" by auto have "(t \ s) \ s" by auto have ii:"finite (t \ s)" using assms(1) by auto have i: "(t \ s) \ (s - (t \ s)) = ( s)" using assms by auto then have "x islimpt s - (t \ s) = x islimpt s" using islimpt_Un_finite[OF ii,where ?t = "s - (t \ s)"] i assms(1) by force then show ?thesis using iii by auto qed lemma ivl_limpt_diff: assumes "finite s" "a < b" "(x::real) \ {a..b} - s" shows "x islimpt {a..b} - s" proof- have "x islimpt {a..b}" proof (cases "x \{a,b}") have i: "finite {a,b}" and ii: "{a, b} \ {a<..{a,b}" then show ?thesis using islimpt_Un_finite[OF i, where ?t= "{a<..{a,b}" then show "x islimpt {a..b}" using assms by auto qed then show "x islimpt {a..b} - s" using islimpt_diff_finite[OF assms(1)] assms by fastforce qed lemma ivl_closure_diff_del: assumes "finite s" "a < b" "(x::real) \ {a..b} - s" shows "x \ closure (({a..b} - s) - {x})" using ivl_limpt_diff islimpt_in_closure assms by blast lemma ivl_not_trivial_limit_within: assumes "finite s" "a < b" "(x::real) \ {a..b} - s" shows "at x within {a..b} - s \ bot" using assms ivl_closure_diff_del not_trivial_limit_within by blast lemma vector_derivative_at_within_non_trivial_limit: "at x within s \ bot \ (f has_vector_derivative f') (at x) \ vector_derivative f (at x within s) = f'" using has_vector_derivative_at_within vector_derivative_within by fastforce lemma vector_derivative_at_within_ivl_diff: "finite s \ a < b \ (x::real) \ {a..b} - s \ (f has_vector_derivative f') (at x) \ vector_derivative f (at x within {a..b} - s) = f'" using vector_derivative_at_within_non_trivial_limit ivl_not_trivial_limit_within by fastforce lemma gamma_deriv_at_within_diff: assumes a_leq_b: "a < b" and x_within_bounds: "x \ {a..b} - s" and gamma_differentiable: "\x \ {a .. b} - s. \ differentiable at x" and s_subset: "s \ {a..b}" and finite_s: "finite s" shows "vector_derivative \ (at x within {a..b} - s) = vector_derivative \ (at x)" using vector_derivative_at_within_ivl_diff [of "s" "a" "b" "x" "\" "vector_derivative \ (at x)"] gamma_differentiable x_within_bounds a_leq_b s_subset finite_s by (auto simp add: vector_derivative_works) lemma gamma_deriv_at_within_gen: assumes a_leq_b: "a < b" and x_within_bounds: "x \ s" and s_subset: "s \ {a..b}" and gamma_differentiable: "\x \ s. \ differentiable at x" shows "vector_derivative \ (at x within ({a..b})) = vector_derivative \ (at x)" using Derivative.vector_derivative_at_within_ivl[of "\" "vector_derivative \ (at x)" "x" "a" "b"] gamma_differentiable x_within_bounds a_leq_b s_subset by (auto simp add: vector_derivative_works) lemma derivative_component_fun_component_at_within_gen: assumes gamma_differentiable: "\x \ s. \ differentiable at x" and s_subset: "s \ {0..1}" shows "\x \ s. vector_derivative (\x. \ x) (at x within {0..1}) \ (i::'a:: euclidean_space) = vector_derivative (\x. \ x \ i) (at x within {0..1})" proof - have gamma_i_component_smooth: "\x \ s. (\x. \ x \ i) differentiable at x" using gamma_differentiable by auto show "\x \ s. vector_derivative (\x. \ x) (at x within {0..1}) \ i = vector_derivative (\x. \ x \ i) (at x within {0..1})" proof fix x::real assume x_within_bounds: "x \ s" have gamma_deriv_at_within: "vector_derivative (\x. \ x) (at x within {0..1}) = vector_derivative (\x. \ x) (at x)" using gamma_deriv_at_within_gen[of "0" "1"] x_within_bounds gamma_differentiable s_subset by (auto simp add: vector_derivative_works) then have gamma_component_deriv_at_within: "vector_derivative (\x. \ x \ i) (at x) = vector_derivative (\x. \ x \ i) (at x within {0..1})" using gamma_deriv_at_within_gen[of "0" "1", where \ = "(\x. \ x \ i)"] x_within_bounds gamma_i_component_smooth s_subset by (auto simp add: vector_derivative_works) have gamma_component_deriv_eq_gamma_deriv_component: "vector_derivative (\x. \ x \ i) (at x) = vector_derivative (\x. \ x) (at x) \ i" using derivative_component_fun_component[of "\" "x" "i"] gamma_differentiable x_within_bounds by auto show "vector_derivative \ (at x within {0..1}) \ i = vector_derivative (\x. \ x \ i) (at x within {0..1})" using gamma_component_deriv_eq_gamma_deriv_component gamma_component_deriv_at_within gamma_deriv_at_within by auto qed qed lemma derivative_component_fun_component_at_within: assumes gamma_differentiable: "\x \ {0 .. 1}. \ differentiable at x" shows "\x \ {0..1}. vector_derivative (\x. \ x) (at x within {0..1}) \ (i::'a:: euclidean_space) = vector_derivative (\x. \ x \ i) (at x within {0..1})" proof - have gamma_i_component_smooth: "\x \ {0 .. 1}. (\x. \ x \ i) differentiable at x" using gamma_differentiable by auto show "\x \ {0..1}. vector_derivative (\x. \ x) (at x within {0..1}) \ i = vector_derivative (\x. \ x \ i) (at x within {0..1})" proof fix x::real assume x_within_bounds: "x \ {0..1}" have gamma_deriv_at_within: "vector_derivative (\x. \ x) (at x within {0..1}) = vector_derivative (\x. \ x) (at x)" using gamma_deriv_at_within[of "0" "1"] x_within_bounds gamma_differentiable by (auto simp add: vector_derivative_works) have gamma_component_deriv_at_within: "vector_derivative (\x. \ x \ i) (at x) = vector_derivative (\x. \ x \ i) (at x within {0..1})" using Derivative.vector_derivative_at_within_ivl[of "(\x. (\ x) \ i)" "vector_derivative (\x. (\ x) \ i) (at x)" "x" "0" "1"] has_vector_derivative_at_within[of "(\x. \ x \ i)" "vector_derivative (\x. \ x \ i) (at x)" "x" "{0..1}"] gamma_i_component_smooth x_within_bounds by (simp add: vector_derivative_works) have gamma_component_deriv_eq_gamma_deriv_component: "vector_derivative (\x. \ x \ i) (at x) = vector_derivative (\x. \ x) (at x) \ i" using derivative_component_fun_component[of "\" "x" "i"] gamma_differentiable x_within_bounds by auto show "vector_derivative \ (at x within {0..1}) \ i = vector_derivative (\x. \ x \ i) (at x within {0..1})" using gamma_component_deriv_eq_gamma_deriv_component gamma_component_deriv_at_within gamma_deriv_at_within by auto qed qed lemma straight_path_diffrentiable_x: fixes b :: "real" and y1 ::"real" assumes gamma_def: "\ = (\x. (b, y2 + y1 * x))" shows "\x. \ differentiable at x" unfolding gamma_def differentiable_def by (fast intro!: derivative_intros) lemma straight_path_diffrentiable_y: fixes b :: "real" and y1 y2 ::"real" assumes gamma_def: "\ = (\x. (y2 + y1 * x , b))" shows "\x. \ differentiable at x" unfolding gamma_def differentiable_def by (fast intro!: derivative_intros) lemma piecewise_C1_differentiable_on_imp_continuous_on: assumes "f piecewise_C1_differentiable_on s" shows "continuous_on s f" using assms by (auto simp add: piecewise_C1_differentiable_on_def) lemma boring_lemma1: fixes f :: "real\real" assumes "(f has_vector_derivative D) (at x)" shows "((\x. (f x, 0)) has_vector_derivative ((D,0::real))) (at x)" proof- have *: "((\x. (f x) *\<^sub>R (1,0)) has_vector_derivative (D *\<^sub>R (1,0))) (at x)" using bounded_linear.has_vector_derivative[OF Real_Vector_Spaces.bounded_linear_scaleR_left assms(1), of "(1,0)"] by auto have "((\x. (f x) *\<^sub>R (1,0)) has_vector_derivative (D,0)) (at x)" proof - have "(D, 0::'a) = D *\<^sub>R (1, 0)" by simp then show ?thesis by (metis (no_types) *) qed then show ?thesis by auto qed lemma boring_lemma2: fixes f :: "real\real" assumes "(f has_vector_derivative D) (at x)" shows "((\x. (0, f x)) has_vector_derivative (0, D)) (at x)" proof- have *: "((\x. (f x) *\<^sub>R (0,1)) has_vector_derivative (D *\<^sub>R (0,1))) (at x)" using bounded_linear.has_vector_derivative[OF Real_Vector_Spaces.bounded_linear_scaleR_left assms(1), of "(0,1)"] by auto then have "((\x. (f x) *\<^sub>R (0,1)) has_vector_derivative ((0,D))) (at x)" using scaleR_Pair Real_Vector_Spaces.real_scaleR_def proof - have "(0::'b, D) = D *\<^sub>R (0, 1)" by auto then show ?thesis by (metis (no_types) *) qed then show ?thesis by auto qed lemma pair_prod_smooth_pw_smooth: assumes "(f::real\real) C1_differentiable_on s" "(g::real\real) piecewise_C1_differentiable_on s" shows "(\x. (f x, g x)) piecewise_C1_differentiable_on s" proof - have f_cont: "continuous_on s f" using assms(1) C1_differentiable_imp_continuous_on by fastforce have g_cont: "continuous_on s g" using assms(2) by (auto simp add: piecewise_C1_differentiable_on_def) obtain t where t:"(finite t \ g C1_differentiable_on s - t)" using assms(2) piecewise_C1_differentiable_on_def by auto show ?thesis using piecewise_C1_differentiable_on_components[where ?f = "(\x. (f x, g x))"] apply (simp add: real_pair_basis) using assms(2) piecewise_C1_differentiable_on_def C1_differentiable_on_subset[OF assms(1) Diff_subset, where ?B1 ="t"] t f_cont g_cont by fastforce qed lemma scale_shift_smooth: shows "(\x. a + b * x) C1_differentiable_on s" proof - show "(\x. a + b * x) C1_differentiable_on s" - using Cauchy_Integral_Theorem.C1_differentiable_on_mult - Cauchy_Integral_Theorem.C1_differentiable_on_add - Cauchy_Integral_Theorem.C1_differentiable_on_const - Cauchy_Integral_Theorem.C1_differentiable_on_ident - by auto + using C1_differentiable_on_mult C1_differentiable_on_add C1_differentiable_on_const + C1_differentiable_on_ident by auto qed lemma open_diff: assumes "finite (t::'a::t1_space set)" "open (s::'a set)" shows "open (s -t)" using assms proof(induction "t") show "open s \ open (s - {})" by auto next fix x::"'a::t1_space" fix F::"'a::t1_space set" assume step: "finite F " " x \ F" "open s" then have i: "(s - insert x F) = (s - F) - {x}" by auto assume ind_hyp: " (open s \ open (s - F))" show "open (s - insert x F)" apply (simp only: i) using open_delete[of "s -F"] ind_hyp[OF step(3)] by auto qed lemma has_derivative_transform_within: assumes "0 < d" and "x \ s" and "\x'\s. dist x' x < d \ f x' = g x'" and "(f has_derivative f') (at x within s)" shows "(g has_derivative f') (at x within s)" using assms unfolding has_derivative_within by (force simp add: intro: Lim_transform_within) lemma has_derivative_transform_within_ivl: assumes "(0::real) < b" and "\x\{a..b} -s. f x = g x" and "x \ {a..b} -s" and "(f has_derivative f') (at x within {a..b} - s)" shows "(g has_derivative f') (at x within {a..b} - s)" using has_derivative_transform_within[of "b" "x" "{a..b} - s"] assms by auto lemma has_vector_derivative_transform_within_ivl: assumes "(0::real) < b" and "\x\{a..b} -s . f x = g x" and "x \ {a..b} - s" and "(f has_vector_derivative f') (at x within {a..b} - s)" shows "(g has_vector_derivative f') (at x within {a..b} - s)" using assms has_derivative_transform_within_ivl apply (auto simp add: has_vector_derivative_def) by blast lemma has_derivative_transform_at: assumes "0 < d" and "\x'. dist x' x < d \ f x' = g x'" and "(f has_derivative f') (at x)" shows "(g has_derivative f') (at x)" using has_derivative_transform_within [of d x UNIV f g f'] assms by simp lemma has_vector_derivative_transform_at: assumes "0 < d" and "\x'. dist x' x < d \ f x' = g x'" and "(f has_vector_derivative f') (at x)" shows "(g has_vector_derivative f') (at x)" using assms unfolding has_vector_derivative_def by (rule has_derivative_transform_at) lemma C1_diff_components_2: assumes "b \ Basis" assumes "f C1_differentiable_on s" shows "(\x. f x \ b) C1_differentiable_on s" proof - obtain D where D:"(\x\s. (f has_derivative (\z. z *\<^sub>R D x)) (at x))" "continuous_on s D" using assms(2) by (fastforce simp add: C1_differentiable_on_def has_vector_derivative_def) show ?thesis proof (simp add: C1_differentiable_on_def has_vector_derivative_def, intro exI conjI) show "continuous_on s (\x. D x \ b)" using D continuous_on_componentwise assms(1) by fastforce show "(\x\s. ((\x. f x \ b) has_derivative (\y. y * (\x. D x \ b) x)) (at x))" using has_derivative_inner_left D(1) by fastforce qed qed lemma eq_smooth: assumes "0 < d" "\x\s. \y. dist x y < d \ f y = g y" (*This crappy condition cannot be loosened :( *) "f C1_differentiable_on s" shows "g C1_differentiable_on s" proof (simp add: C1_differentiable_on_def) obtain D where D: "(\x\s. (f has_vector_derivative D x) (at x)) \ continuous_on s D" using assms by (auto simp add: C1_differentiable_on_def) then have f: "(\x\s. (g has_vector_derivative D x) (at x))" using assms(1-2) by (metis dist_commute has_vector_derivative_transform_at) have "(\x\s. (g has_vector_derivative D x) (at x)) \ continuous_on s D" using D f by auto then show "\D. (\x\s. (g has_vector_derivative D x) (at x)) \ continuous_on s D" by metis qed lemma eq_pw_smooth: assumes "0 < d" "\x\s. \y. dist x y < d \ f y = g y" (*This crappy condition cannot be loosened :( *) "\x\s. f x = g x" "f piecewise_C1_differentiable_on s" shows "g piecewise_C1_differentiable_on s" proof (simp add: piecewise_C1_differentiable_on_def) have g_cont: "continuous_on s g" using assms piecewise_C1_differentiable_const by (simp add: piecewise_C1_differentiable_on_def) obtain t where t: "finite t \ f C1_differentiable_on s - t" using assms by (auto simp add: piecewise_C1_differentiable_on_def) then have "g C1_differentiable_on s - t" using assms eq_smooth by blast then show "continuous_on s g \ (\t. finite t \ g C1_differentiable_on s - t)" using t g_cont by metis qed lemma scale_piecewise_C1_differentiable_on: assumes "f piecewise_C1_differentiable_on s" shows "(\x. (c::real) * (f x)) piecewise_C1_differentiable_on s" proof (simp add: piecewise_C1_differentiable_on_def, intro conjI) show "continuous_on s (\x. c * f x)" using assms continuous_on_mult_left by (auto simp add: piecewise_C1_differentiable_on_def) show "\t. finite t \ (\x. c * f x) C1_differentiable_on s - t" using assms continuous_on_mult_left by (auto simp add: piecewise_C1_differentiable_on_def) qed lemma eq_smooth_gen: assumes "f C1_differentiable_on s" "\x. f x = g x" shows "g C1_differentiable_on s" using assms unfolding C1_differentiable_on_def by (metis (no_types, lifting) has_vector_derivative_weaken UNIV_I top_greatest) lemma subpath_compose: shows "(subpath a b \) = \ o (\x. (b - a) * x + a)" by (auto simp add: subpath_def) lemma subpath_smooth: assumes "\ C1_differentiable_on {0..1}" "0 \ a" "a < b" "b \ 1" shows "(subpath a b \) C1_differentiable_on {0..1}" proof- have "\ C1_differentiable_on {a..b}" apply (rule C1_differentiable_on_subset) using assms by auto then have "\ C1_differentiable_on (\x. (b - a) * x + a) ` {0..1}" using \a < b\ closed_segment_eq_real_ivl closed_segment_real_eq by auto moreover have "finite ({0..1} \ (\x. (b - a) * x + a) -` {x})" for x proof - have "((\x. (b - a) * x + a) -` {x}) = {(x -a) /(b-a)}" using assms by (auto simp add: divide_simps) then show ?thesis by auto qed ultimately show ?thesis by (force simp add: subpath_compose intro: C1_differentiable_compose derivative_intros) qed lemma has_vector_derivative_divide[derivative_intros]: fixes a :: "'a::real_normed_field" shows "(f has_vector_derivative x) F \ ((\x. f x / a) has_vector_derivative (x / a)) F" unfolding divide_inverse by (fact has_vector_derivative_mult_left) end diff --git a/thys/Green/Green.thy b/thys/Green/Green.thy --- a/thys/Green/Green.thy +++ b/thys/Green/Green.thy @@ -1,3117 +1,3117 @@ theory Green imports Paths Derivs Integrals General_Utils begin lemma frontier_Un_subset_Un_frontier: "frontier (s \ t) \ (frontier s) \ (frontier t)" by (simp add: frontier_def Un_Diff) (auto simp add: closure_def interior_def islimpt_def) definition has_partial_derivative:: "(('a::euclidean_space) \ 'b::euclidean_space) \ 'a \ ('a \ 'b) \ ('a) \ bool" where "has_partial_derivative F base_vec F' a \ ((\x::'a::euclidean_space. F( (a - ((a \ base_vec) *\<^sub>R base_vec)) + (x \ base_vec) *\<^sub>R base_vec )) has_derivative F') (at a)" definition has_partial_vector_derivative:: "(('a::euclidean_space) \ 'b::euclidean_space) \ 'a \ ( 'b) \ ('a) \ bool" where "has_partial_vector_derivative F base_vec F' a \ ((\x. F( (a - ((a \ base_vec) *\<^sub>R base_vec)) + x *\<^sub>R base_vec )) has_vector_derivative F') (at (a \ base_vec))" definition partially_vector_differentiable where "partially_vector_differentiable F base_vec p \ (\F'. has_partial_vector_derivative F base_vec F' p)" definition partial_vector_derivative:: "(('a::euclidean_space) \ 'b::euclidean_space) \ 'a \ 'a \ 'b" where "partial_vector_derivative F base_vec a \ (vector_derivative (\x. F( (a - ((a \ base_vec) *\<^sub>R base_vec)) + x *\<^sub>R base_vec)) (at (a \ base_vec)))" lemma partial_vector_derivative_works: assumes "partially_vector_differentiable F base_vec a" shows "has_partial_vector_derivative F base_vec (partial_vector_derivative F base_vec a) a" proof - obtain F' where F'_prop: "((\x. F( (a - ((a \ base_vec) *\<^sub>R base_vec)) + x *\<^sub>R base_vec )) has_vector_derivative F') (at (a \ base_vec))" using assms partially_vector_differentiable_def has_partial_vector_derivative_def by blast show ?thesis using Derivative.differentiableI_vector[OF F'_prop] by(simp add: vector_derivative_works partial_vector_derivative_def[symmetric] has_partial_vector_derivative_def[symmetric]) qed lemma fundamental_theorem_of_calculus_partial_vector: fixes a b:: "real" and F:: "('a::euclidean_space \ 'b::euclidean_space)" and i:: "'a" and j:: "'b" and F_j_i:: "('a::euclidean_space \ real)" assumes a_leq_b: "a \ b" and Base_vecs: "i \ Basis" "j \ Basis" and no_i_component: "c \ i = 0 " and has_partial_deriv: "\p \ D. has_partial_vector_derivative (\x. (F x) \ j) i (F_j_i p) p" and domain_subset_of_D: "{x *\<^sub>R i + c |x. a \ x \ x \ b} \ D" shows "((\x. F_j_i( x *\<^sub>R i + c)) has_integral F(b *\<^sub>R i + c) \ j - F(a *\<^sub>R i + c) \ j) (cbox a b)" proof - let ?domain = "{v. \x. a \ x \ x \ b \ v = x *\<^sub>R i + c}" have "\x\ ?domain. has_partial_vector_derivative (\x. (F x) \ j) i (F_j_i x) x" using has_partial_deriv domain_subset_of_D by auto then have "\x\ (cbox a b). ((\x. F(x *\<^sub>R i + c) \ j) has_vector_derivative (F_j_i( x *\<^sub>R i + c))) (at x)" proof(clarsimp simp add: has_partial_vector_derivative_def) fix x::real assume range_of_x: "a \ x" "x \ b" assume ass2: "\x. (\z\a. z \ b \ x = z *\<^sub>R i + c) \ ((\z. F(x - (x \ i) *\<^sub>R i + z *\<^sub>R i) \ j) has_vector_derivative F_j_i x) (at (x \ i))" have "((\z. F((x *\<^sub>R i + c) - ((x *\<^sub>R i + c) \ i) *\<^sub>R i + z *\<^sub>R i) \ j) has_vector_derivative F_j_i (x *\<^sub>R i + c)) (at ((x *\<^sub>R i + c) \ i)) " using range_of_x ass2 by auto then have 0: "((\x. F( c + x *\<^sub>R i) \ j) has_vector_derivative F_j_i (x *\<^sub>R i + c)) (at x)" by (simp add: assms(2) inner_left_distrib no_i_component) have 1: "(\x. F(x *\<^sub>R i + c) \ j) = (\x. F(c + x *\<^sub>R i) \ j)" by (simp add: add.commute) then show "((\x. F(x *\<^sub>R i + c) \ j) has_vector_derivative F_j_i (x *\<^sub>R i + c)) (at x)" using 0 and 1 by auto qed then have "\x\ (cbox a b). ((\x. F(x *\<^sub>R i + c) \ j) has_vector_derivative (F_j_i( x *\<^sub>R i + c))) (at_within x (cbox a b))" using has_vector_derivative_at_within by blast then show "( (\x. F_j_i( x *\<^sub>R i + c)) has_integral F(b *\<^sub>R i + c) \ j - F(a *\<^sub>R i + c) \ j) (cbox a b)" using fundamental_theorem_of_calculus[of "a" "b" "(\x::real. F(x *\<^sub>R i + c) \ j)" "(\x::real. F_j_i( x *\<^sub>R i + c))"] and a_leq_b by auto qed lemma fundamental_theorem_of_calculus_partial_vector_gen: fixes k1 k2:: "real" and F:: "('a::euclidean_space \ 'b::euclidean_space)" and i:: "'a" and F_i:: "('a::euclidean_space \ 'b)" assumes a_leq_b: "k1 \ k2" and unit_len: "i \ i = 1" and no_i_component: "c \ i = 0 " and has_partial_deriv: "\p \ D. has_partial_vector_derivative F i (F_i p) p" and domain_subset_of_D: "{v. \x. k1 \ x \ x \ k2 \ v = x *\<^sub>R i + c} \ D" shows "( (\x. F_i( x *\<^sub>R i + c)) has_integral F(k2 *\<^sub>R i + c) - F(k1 *\<^sub>R i + c)) (cbox k1 k2)" proof - let ?domain = "{v. \x. k1 \ x \ x \ k2 \ v = x *\<^sub>R i + c}" have "\x\ ?domain. has_partial_vector_derivative F i (F_i x) x" using has_partial_deriv domain_subset_of_D by auto then have "\x\ (cbox k1 k2). ((\x. F(x *\<^sub>R i + c)) has_vector_derivative (F_i( x *\<^sub>R i + c))) (at x)" proof (clarsimp simp add: has_partial_vector_derivative_def) fix x::real assume range_of_x: "k1 \ x" "x \ k2" assume ass2: "\x. (\z\k1. z \ k2 \ x = z *\<^sub>R i + c) \ ((\z. F(x - (x \ i) *\<^sub>R i + z *\<^sub>R i)) has_vector_derivative F_i x) (at (x \ i))" have "((\z. F((x *\<^sub>R i + c) - ((x *\<^sub>R i + c) \ i) *\<^sub>R i + z *\<^sub>R i)) has_vector_derivative F_i (x *\<^sub>R i + c)) (at ((x *\<^sub>R i + c) \ i)) " using range_of_x ass2 by auto then have 0: "((\x. F( c + x *\<^sub>R i)) has_vector_derivative F_i (x *\<^sub>R i + c)) (at x)" by (simp add: inner_commute inner_right_distrib no_i_component unit_len) have 1: "(\x. F(x *\<^sub>R i + c)) = (\x. F(c + x *\<^sub>R i))" by (simp add: add.commute) then show "((\x. F(x *\<^sub>R i + c)) has_vector_derivative F_i (x *\<^sub>R i + c)) (at x)" using 0 and 1 by auto qed then have "\x\ (cbox k1 k2). ((\x. F(x *\<^sub>R i + c) ) has_vector_derivative (F_i( x *\<^sub>R i + c))) (at_within x (cbox k1 k2))" using has_vector_derivative_at_within by blast then show "( (\x. F_i( x *\<^sub>R i + c)) has_integral F(k2 *\<^sub>R i + c) - F(k1 *\<^sub>R i + c)) (cbox k1 k2)" using fundamental_theorem_of_calculus[of "k1" "k2" "(\x::real. F(x *\<^sub>R i + c))" "(\x::real. F_i( x *\<^sub>R i + c))"] and a_leq_b by auto qed lemma add_scale_img: assumes "a < b" shows "(\x::real. a + (b - a) * x) ` {0 .. 1} = {a .. b}" using assms apply (auto simp: algebra_simps affine_ineq image_iff) using less_eq_real_def apply force apply (rule_tac x="(x-a)/(b-a)" in bexI) apply (auto simp: field_simps) done lemma add_scale_img': assumes "a \ b" shows "(\x::real. a + (b - a) * x) ` {0 .. 1} = {a .. b}" proof (cases "a = b") case True then show ?thesis by force next case False then show ?thesis using add_scale_img assms by auto qed definition analytically_valid:: "'a::euclidean_space set \ ('a \ 'b::{euclidean_space,times,one}) \ 'a \ bool" where "analytically_valid s F i \ (\a \ s. partially_vector_differentiable F i a) \ continuous_on s F \ \ \TODO: should we replace this with saying that \F\ is partially diffrerentiable on \Dy\,\ \ \i.e. there is a partial derivative on every dimension\ integrable lborel (\p. (partial_vector_derivative F i) p * indicator s p) \ (\x. integral UNIV (\y. (partial_vector_derivative F i (y *\<^sub>R i + x *\<^sub>R (\ b \(Basis - {i}). b))) * (indicator s (y *\<^sub>R i + x *\<^sub>R (\b \ Basis - {i}. b)) ))) \ borel_measurable lborel" (*(\x. integral UNIV (\y. ((partial_vector_derivative F i) (y, x)) * (indicator s (y, x)))) \ borel_measurable lborel)"*) lemma analytically_valid_imp_part_deriv_integrable_on: assumes "analytically_valid (s::(real*real) set) (f::(real*real)\ real) i" shows "(partial_vector_derivative f i) integrable_on s" proof - have "integrable lborel (\p. partial_vector_derivative f i p * indicator s p)" using assms by(simp add: analytically_valid_def indic_ident) then have "integrable lborel (\p::(real*real). if p \ s then partial_vector_derivative f i p else 0)" using indic_ident[of "partial_vector_derivative f i"] by (simp add: indic_ident) then have "(\x. if x \ s then partial_vector_derivative f i x else 0) integrable_on UNIV" using Equivalence_Lebesgue_Henstock_Integration.integrable_on_lborel by auto then show "(partial_vector_derivative f i) integrable_on s" using integrable_restrict_UNIV by auto qed (*******************************************************************************) definition typeII_twoCube :: "((real * real) \ (real * real)) \ bool" where "typeII_twoCube twoC \ \a b g1 g2. a < b \ (\x \ {a..b}. g2 x \ g1 x) \ twoC = (\(y, x). ((1 - y) * (g2 ((1-x)*a + x*b)) + y * (g1 ((1-x)*a + x*b)), (1-x)*a + x*b)) \ g1 piecewise_C1_differentiable_on {a .. b} \ g2 piecewise_C1_differentiable_on {a .. b}" abbreviation unit_cube where "unit_cube \ cbox (0,0) (1::real,1::real)" definition cubeImage:: "two_cube \ ((real*real) set)" where "cubeImage twoC \ (twoC ` unit_cube)" lemma typeII_twoCubeImg: assumes "typeII_twoCube twoC" shows "\a b g1 g2. a < b \ (\x \ {a .. b}. g2 x \ g1 x) \ cubeImage twoC = {(y,x). x \ {a..b} \ y \ {g2 x .. g1 x}} \ twoC = (\(y, x). ((1 - y) * g2 ((1 - x) * a + x * b) + y * g1 ((1 - x) * a + x * b), (1 - x) * a + x * b)) \ g1 piecewise_C1_differentiable_on {a .. b} \ g2 piecewise_C1_differentiable_on {a .. b} " using assms proof (simp add: typeII_twoCube_def cubeImage_def image_def) assume " \a b. a < b \ (\g1 g2. (\x\{a..b}. g2 x \ g1 x) \ twoC = (\(y, x). ((1 - y) * g2 ((1 - x) * a + x * b) + y * g1 ((1 - x) * a + x * b), (1 - x) * a + x * b)) \ g1 piecewise_C1_differentiable_on {a..b} \ g2 piecewise_C1_differentiable_on {a..b})" then obtain a b g1 g2 where twoCisTypeII: "a < b" "(\x\{a..b}. g2 x \ g1 x)" "twoC = (\(y, x). ((1 - y) * g2 ((1 - x) * a + x * b) + y * g1 ((1 - x) * a + x * b), (1 - x) * a + x * b))" "g1 piecewise_C1_differentiable_on {a .. b}" "g2 piecewise_C1_differentiable_on {a .. b}" by auto have ab1: "a - x * a + x * b \ b" if a1: "0 \ x" "x \ 1" for x using that apply (simp add: algebra_simps) by (metis affine_ineq less_eq_real_def mult.commute twoCisTypeII(1)) have ex: "\z\Green.unit_cube. (d, c) = (case z of (y, x) \ (g2 (a - x * a + x * b) - y * g2 (a - x * a + x * b) + y * g1 (a - x * a + x * b), a - x * a + x * b))" if c_bounds: "a \ c" "c \ b" and d_bounds: "g2 c \ d" "d \ g1 c" for c d proof - have b_minus_a_nzero: "b - a \ 0" using twoCisTypeII(1) by auto have x_witness: "\k1. c = (1 - k1)*a + k1 * b \ 0 \ k1 \ k1 \ 1" apply (rule_tac x="(c - a)/(b - a)" in exI) using c_bounds \a < b\ apply (simp add: divide_simps algebra_simps) using sum_sqs_eq by blast have y_witness: "\k2. d = (1 - k2)*(g2 c) + k2 * (g1 c) \ 0 \ k2 \ k2 \ 1" proof (cases "g1 c - g2 c = 0") case True with d_bounds show ?thesis by (fastforce simp add: algebra_simps) next case False let ?k2 = "(d - g2 c)/(g1 c - g2 c)" have k2_in_bounds: "0 \ ?k2 \ ?k2 \ 1" using twoCisTypeII(2) c_bounds d_bounds False by simp have "d = (1 - ?k2) * g2 c + ?k2 * g1 c" using False sum_sqs_eq by (fastforce simp add: divide_simps algebra_simps) with k2_in_bounds show ?thesis by fastforce qed show "\x\unit_cube. (d, c) = (case x of (y, x) \ (g2 (a - x * a + x * b) - y * g2 (a - x * a + x * b) + y * g1 (a - x * a + x * b), a - x * a + x * b))" using x_witness y_witness by (force simp add: left_diff_distrib) qed have "{y. \x\unit_cube. y = twoC x} = {(y, x). a \ x \ x \ b \ g2 x \ y \ y \ g1 x}" apply (auto simp add: twoCisTypeII ab1 left_diff_distrib ex) using order.order_iff_strict twoCisTypeII(1) apply fastforce apply (smt affine_ineq atLeastAtMost_iff mult_left_mono twoCisTypeII)+ done then show "\a b. a < b \ (\g1 g2. (\x\{a..b}. g2 x \ g1 x) \ {y. \x\unit_cube. y = twoC x} = {(y, x). a \ x \ x \ b \ g2 x \ y \ y \ g1 x} \ twoC = (\(y, x). ((1 - y) * g2 ((1 - x) * a + x * b) + y * g1 ((1 - x) * a + x * b), (1 - x) * a + x * b)) \ g1 piecewise_C1_differentiable_on {a..b} \ g2 piecewise_C1_differentiable_on {a..b})" using twoCisTypeII by blast qed definition horizontal_boundary :: "two_cube \ one_chain" where "horizontal_boundary twoC \ {(1, (\x. twoC(x,0))), (-1, (\x. twoC(x,1)))}" definition vertical_boundary :: "two_cube \ one_chain" where "vertical_boundary twoC \ {(-1, (\y. twoC(0,y))), (1, (\y. twoC(1,y)))}" definition boundary :: "two_cube \ one_chain" where "boundary twoC \ horizontal_boundary twoC \ vertical_boundary twoC" definition valid_two_cube where "valid_two_cube twoC \ card (boundary twoC) = 4" definition two_chain_integral:: "two_chain \ ((real*real)\(real)) \ real" where "two_chain_integral twoChain F \ \C\twoChain. (integral (cubeImage C) F)" definition valid_two_chain where "valid_two_chain twoChain \ (\twoCube \ twoChain. valid_two_cube twoCube) \ pairwise (\c1 c2. ((boundary c1) \ (boundary c2)) = {}) twoChain \ inj_on cubeImage twoChain" definition two_chain_boundary:: "two_chain \ one_chain" where "two_chain_boundary twoChain == \(boundary ` twoChain)" definition gen_division where "gen_division s S \ (finite S \ (\S = s) \ pairwise (\X Y. negligible (X \ Y)) S)" definition two_chain_horizontal_boundary:: "two_chain \ one_chain" where "two_chain_horizontal_boundary twoChain \ \(horizontal_boundary ` twoChain)" definition two_chain_vertical_boundary:: "two_chain \ one_chain" where "two_chain_vertical_boundary twoChain \ \(vertical_boundary ` twoChain)" definition only_horizontal_division where "only_horizontal_division one_chain two_chain \ \\ \. finite \ \ finite \ \ (\(k,\) \ \. (\(k', \') \ two_chain_horizontal_boundary two_chain. (\a \ {0..1}. \b \ {0..1}. a \ b \ subpath a b \' = \))) \ (common_sudiv_exists (two_chain_vertical_boundary two_chain) \ \ common_reparam_exists \ (two_chain_vertical_boundary two_chain)) \ boundary_chain \ \ one_chain = \ \ \ \ (\(k,\)\\. valid_path \)" lemma sum_zero_set: assumes "\x \ s. f x = 0" "finite s" "finite t" shows "sum f (s \ t) = sum f t" using assms by (simp add: IntE sum.union_inter_neutral sup_commute) abbreviation "valid_typeII_division s twoChain \ ((\twoCube \ twoChain. typeII_twoCube twoCube) \ (gen_division s (cubeImage ` twoChain)) \ (valid_two_chain twoChain))" lemma two_chain_vertical_boundary_is_boundary_chain: shows "boundary_chain (two_chain_vertical_boundary twoChain)" by(simp add: boundary_chain_def two_chain_vertical_boundary_def vertical_boundary_def) lemma two_chain_horizontal_boundary_is_boundary_chain: shows "boundary_chain (two_chain_horizontal_boundary twoChain)" by(simp add: boundary_chain_def two_chain_horizontal_boundary_def horizontal_boundary_def) definition typeI_twoCube :: "two_cube \ bool" where "typeI_twoCube (twoC::two_cube) \ \a b g1 g2. a < b \ (\x \ {a..b}. g2 x \ g1 x) \ twoC = (\(x,y). ((1-x)*a + x*b, (1 - y) * (g2 ((1-x)*a + x*b)) + y * (g1 ((1-x)*a + x*b)))) \ g1 piecewise_C1_differentiable_on {a..b} \ g2 piecewise_C1_differentiable_on {a..b}" lemma typeI_twoCubeImg: assumes "typeI_twoCube twoC" shows "\a b g1 g2. a < b \ (\x \ {a .. b}. g2 x \ g1 x) \ cubeImage twoC = {(x,y). x \ {a..b} \ y \ {g2 x .. g1 x}} \ twoC = (\(x, y). ((1 - x) * a + x * b, (1 - y) * g2 ((1 - x) * a + x * b) + y * g1 ((1 - x) * a + x * b))) \ g1 piecewise_C1_differentiable_on {a .. b} \ g2 piecewise_C1_differentiable_on {a .. b} " proof - have "\a b. a < b \ (\g1 g2. (\x\{a..b}. g2 x \ g1 x) \ twoC = (\(x, y). ((1 - x) * a + x * b, (1 - y) * g2 ((1 - x) * a + x * b) + y * g1 ((1 - x) * a + x * b)))\ g1 piecewise_C1_differentiable_on {a .. b} \ g2 piecewise_C1_differentiable_on {a .. b})" using assms by (simp add: typeI_twoCube_def) then obtain a b g1 g2 where twoCisTypeI: "a < b" "(\x\{a..b}. g2 x \ g1 x)" "twoC = (\(x, y). ((1 - x) * a + x * b, (1 - y) * g2 ((1 - x) * a + x * b) + y * g1 ((1 - x) * a + x * b)))" "g1 piecewise_C1_differentiable_on {a .. b}" "g2 piecewise_C1_differentiable_on {a .. b}" by auto have ex: "\z\Green.unit_cube. (c, d) = (case z of (x, y) \ (a - x * a + x * b, g2 (a - x * a + x * b) - y * g2 (a - x * a + x * b) + y * g1 (a - x * a + x * b)))" if c_bounds: "a \ c" "c \ b" and d_bounds: "g2 c \ d" "d \ g1 c" for c d proof - have x_witness: "\k1. c = (1 - k1)*a + k1 * b \ 0 \ k1 \ k1 \ 1" proof - let ?k1 = "(c - a)/(b - a)" have k1_in_bounds: "0 \ (c - a)/(b - a) \ (c - a)/(b - a) \ 1" using twoCisTypeI(1) c_bounds by simp have "c = (1 - ?k1)*a + ?k1 * b" using twoCisTypeI(1) sum_sqs_eq by (auto simp add: divide_simps algebra_simps) then show ?thesis using twoCisTypeI k1_in_bounds by fastforce qed have y_witness: "\k2. d = (1 - k2)*(g2 c) + k2 * (g1 c) \ 0 \ k2 \ k2 \ 1" proof (cases "g1 c - g2 c = 0") case True with d_bounds show ?thesis by force next case False let ?k2 = "(d - g2 c)/(g1 c - g2 c)" have k2_in_bounds: "0 \ ?k2 \ ?k2 \ 1" using twoCisTypeI(2) c_bounds d_bounds False by simp have "d = (1 - ?k2) * g2 c + ?k2 * g1 c" using False apply (simp add: divide_simps algebra_simps) using sum_sqs_eq by fastforce then show ?thesis using k2_in_bounds by fastforce qed show "\x\unit_cube. (c, d) = (case x of (x, y) \ (a - x * a + x * b, g2 (a - x * a + x * b) - y * g2 (a - x * a + x * b) + y * g1 (a - x * a + x * b)))" using x_witness y_witness by (force simp add: left_diff_distrib) qed have "{y. \x\unit_cube. y = twoC x} = {(x, y). a \ x \ x \ b \ g2 x \ y \ y \ g1 x}" apply (auto simp add: twoCisTypeI left_diff_distrib ex) using less_eq_real_def twoCisTypeI(1) apply auto[1] apply (smt affine_ineq twoCisTypeI) apply (smt affine_ineq atLeastAtMost_iff mult_left_mono twoCisTypeI)+ done then show ?thesis unfolding cubeImage_def image_def using twoCisTypeI by auto qed lemma typeI_cube_explicit_spec: assumes "typeI_twoCube twoC" shows "\a b g1 g2. a < b \ (\x \ {a .. b}. g2 x \ g1 x) \ cubeImage twoC = {(x,y). x \ {a..b} \ y \ {g2 x .. g1 x}} \ twoC = (\(x, y). ((1 - x) * a + x * b, (1 - y) * g2 ((1 - x) * a + x * b) + y * g1 ((1 - x) * a + x * b))) \ g1 piecewise_C1_differentiable_on {a .. b} \ g2 piecewise_C1_differentiable_on {a .. b} \ (\x. twoC(x, 0)) = (\x. (a + (b - a) * x, g2 (a + (b - a) * x))) \ (\y. twoC(1, y)) = (\x. (b, g2 b + x *\<^sub>R (g1 b - g2 b))) \ (\x. twoC(x, 1)) = (\x. (a + (b - a) * x, g1 (a + (b - a) * x))) \ (\y. twoC(0, y)) = (\x. (a, g2 a + x *\<^sub>R (g1 a - g2 a)))" proof - let ?bottom_edge = "(\x. twoC(x, 0))" let ?right_edge = "(\y. twoC(1, y))" let ?top_edge = "(\x. twoC(x, 1))" let ?left_edge = "(\y. twoC(0, y))" obtain a b g1 g2 where twoCisTypeI: "a < b" "(\x \ cbox a b. g2 x \ g1 x)" "cubeImage twoC = {(x,y). x \ cbox a b \ y \ cbox (g2 x) (g1 x)}" "twoC = (\(x, y). ((1 - x) * a + x * b, (1 - y) * g2 ((1 - x) * a + x * b) + y * g1 ((1 - x) * a + x * b)))" "g1 piecewise_C1_differentiable_on {a .. b}" "g2 piecewise_C1_differentiable_on {a .. b}" using assms and typeI_twoCubeImg[of"twoC"] by auto have bottom_edge_explicit: "?bottom_edge = (\x. (a + (b - a) * x, g2 (a + (b - a) * x)))" by (simp add: twoCisTypeI(4) algebra_simps) have right_edge_explicit: "?right_edge = (\x. (b, g2 b + x *\<^sub>R (g1 b - g2 b)))" by (simp add: twoCisTypeI(4) algebra_simps) have top_edge_explicit: "?top_edge = (\x. (a + (b - a) * x, g1 (a + (b - a) * x)))" by (simp add: twoCisTypeI(4) algebra_simps) have left_edge_explicit: "?left_edge = (\x. (a, g2 a + x *\<^sub>R (g1 a - g2 a)))" by (simp add: twoCisTypeI(4) algebra_simps) show ?thesis using bottom_edge_explicit right_edge_explicit top_edge_explicit left_edge_explicit twoCisTypeI by auto qed lemma typeI_twoCube_smooth_edges: assumes "typeI_twoCube twoC" "(k,\) \ boundary twoC" shows "\ piecewise_C1_differentiable_on {0..1}" proof - let ?bottom_edge = "(\x. twoC(x, 0))" let ?right_edge = "(\y. twoC(1, y))" let ?top_edge = "(\x. twoC(x, 1))" let ?left_edge = "(\y. twoC(0, y))" obtain a b g1 g2 where twoCisTypeI: "a < b" "(\x \ cbox a b. g2 x \ g1 x)" "cubeImage twoC = {(x,y). x \ cbox a b \ y \ cbox (g2 x) (g1 x)}" "twoC = (\(x, y). ((1 - x) * a + x * b, (1 - y) * g2 ((1 - x) * a + x * b) + y * g1 ((1 - x) * a + x * b)))" "g1 piecewise_C1_differentiable_on {a .. b}" "g2 piecewise_C1_differentiable_on {a .. b}" "(\x. twoC(x, 0)) = (\x. (a + (b - a) * x, g2 (a + (b - a) * x)))" "(\y. twoC(1, y)) = (\x. (b, g2 b + x *\<^sub>R (g1 b - g2 b)))" "(\x. twoC(x, 1)) = (\x. (a + (b - a) * x, g1 (a + (b - a) * x)))" "(\y. twoC(0, y)) = (\x. (a, g2 a + x *\<^sub>R (g1 a - g2 a)))" using assms and typeI_cube_explicit_spec[of"twoC"] by auto have bottom_edge_smooth: "(\x. twoC (x, 0)) piecewise_C1_differentiable_on {0..1}" proof - have "\x. (\x. (a + (b - a) * x))-` {x} = {(x - a)/(b -a)}" using twoCisTypeI(1) by(auto simp add: Set.vimage_def) then have finite_vimg: "\x. finite({0..1} \ (\x. (a + (b - a) * x))-` {x})" by auto have scale_shif_smth: "(\x. (a + (b - a) * x)) C1_differentiable_on {0..1}" using scale_shift_smooth by auto then have scale_shif_pw_smth: "(\x. (a + (b - a) * x)) piecewise_C1_differentiable_on {0..1}" using C1_differentiable_imp_piecewise by blast have g2_smooth: "g2 piecewise_C1_differentiable_on (\x. a + (b - a) * x) ` {0..1}" using add_scale_img[OF twoCisTypeI(1)] twoCisTypeI(6) by auto have "(\x. g2 (a + (b - a) * x)) piecewise_C1_differentiable_on {0..1}" - using Cauchy_Integral_Theorem.piecewise_C1_differentiable_compose[OF scale_shif_pw_smth g2_smooth finite_vimg] + using piecewise_C1_differentiable_compose[OF scale_shif_pw_smth g2_smooth finite_vimg] by (auto simp add: o_def) then have "(\x::real. (a + (b - a) * x, g2 (a + (b - a) * x))) piecewise_C1_differentiable_on {0..1}" using all_components_smooth_one_pw_smooth_is_pw_smooth[where f = "(\x::real. (a + (b - a) * x, g2 (a + (b - a) * x)))"] apply (simp only: real_pair_basis) by fastforce then show ?thesis using twoCisTypeI(7) by auto qed have top_edge_smooth: "?top_edge piecewise_C1_differentiable_on {0..1}" proof - have "\x. (\x. (a + (b - a) * x))-` {x} = {(x - a)/(b -a)}" using twoCisTypeI(1) by(auto simp add: Set.vimage_def) then have finite_vimg: "\x. finite({0..1} \ (\x. (a + (b - a) * x))-` {x})" by auto have scale_shif_smth: "(\x. (a + (b - a) * x)) C1_differentiable_on {0..1}" using scale_shift_smooth by auto then have scale_shif_pw_smth: "(\x. (a + (b - a) * x)) piecewise_C1_differentiable_on {0..1}" using C1_differentiable_imp_piecewise by blast have g1_smooth: "g1 piecewise_C1_differentiable_on (\x. a + (b - a) * x) ` {0..1}" using add_scale_img[OF twoCisTypeI(1)] twoCisTypeI(5) by auto have "(\x. g1 (a + (b - a) * x)) piecewise_C1_differentiable_on {0..1}" - using Cauchy_Integral_Theorem.piecewise_C1_differentiable_compose[OF scale_shif_pw_smth g1_smooth finite_vimg] + using piecewise_C1_differentiable_compose[OF scale_shif_pw_smth g1_smooth finite_vimg] by (auto simp add: o_def) then have "(\x. (a + (b - a) * x, g1 (a + (b - a) * x))) piecewise_C1_differentiable_on {0..1}" using all_components_smooth_one_pw_smooth_is_pw_smooth[where f = "(\x. (a + (b - a) * x, g1 (a + (b - a) * x)))"] apply (simp only: real_pair_basis) by fastforce then show ?thesis using twoCisTypeI(9) by auto qed have right_edge_smooth: "?right_edge piecewise_C1_differentiable_on {0..1}" proof - have "(\x. (g2 b + x *\<^sub>R (g1 b - g2 b))) C1_differentiable_on {0..1}" using scale_shift_smooth C1_differentiable_imp_piecewise by auto then have "(\x. (g2 b + x *\<^sub>R (g1 b - g2 b))) piecewise_C1_differentiable_on {0..1}" using C1_differentiable_imp_piecewise by fastforce then have "(\x. (b, g2 b + x *\<^sub>R (g1 b - g2 b))) piecewise_C1_differentiable_on {0..1}" using pair_prod_smooth_pw_smooth by auto then show ?thesis using twoCisTypeI(8) by auto qed have left_edge_smooth: "?left_edge piecewise_C1_differentiable_on {0..1}" proof - have "(\x. (g2 a + x *\<^sub>R (g1 a - g2 a))) C1_differentiable_on {0..1}" using scale_shift_smooth by auto then have "(\x. (g2 a + x *\<^sub>R (g1 a - g2 a))) piecewise_C1_differentiable_on {0..1}" using C1_differentiable_imp_piecewise by fastforce then have "(\x. (a, g2 a + x *\<^sub>R (g1 a - g2 a))) piecewise_C1_differentiable_on {0..1}" using pair_prod_smooth_pw_smooth by auto then show ?thesis using twoCisTypeI(10) by auto qed have "\ = ?bottom_edge \ \ = ?right_edge \ \ = ?top_edge \ \ = ?left_edge" using assms by (auto simp add: boundary_def horizontal_boundary_def vertical_boundary_def) then show ?thesis using left_edge_smooth right_edge_smooth top_edge_smooth bottom_edge_smooth by auto qed lemma two_chain_integral_eq_integral_divisable: assumes f_integrable: "\twoCube \ twoChain. F integrable_on cubeImage twoCube" and gen_division: "gen_division s (cubeImage ` twoChain)" and valid_two_chain: "valid_two_chain twoChain" shows "integral s F = two_chain_integral twoChain F" proof - show "integral s F = two_chain_integral twoChain F" proof (simp add: two_chain_integral_def) have partial_deriv_integrable: "\twoCube \ twoChain. ((F) has_integral (integral (cubeImage twoCube) (F))) (cubeImage twoCube)" using f_integrable by auto then have partial_deriv_integrable: "\twoCubeImg. twoCubeImg \ cubeImage ` twoChain \ (F has_integral (integral (twoCubeImg) F)) (twoCubeImg)" using Henstock_Kurzweil_Integration.integrable_neg by force have finite_images: "finite (cubeImage ` twoChain)" using gen_division gen_division_def by auto have negligible_images: "pairwise (\S S'. negligible (S \ S')) (cubeImage ` twoChain)" using gen_division by (auto simp add: gen_division_def pairwise_def) have inj: "inj_on cubeImage twoChain" using valid_two_chain by (simp add: inj_on_def valid_two_chain_def) have "integral s F = (\twoCubeImg\cubeImage ` twoChain. integral twoCubeImg F)" using has_integral_Union[OF finite_images partial_deriv_integrable negligible_images] gen_division by (auto simp add: gen_division_def) also have "\ = (\C\twoChain. integral (cubeImage C) F)" using sum.reindex inj by auto finally show "integral s F = (\C\twoChain. integral (cubeImage C) F)" . qed qed definition only_vertical_division where "only_vertical_division one_chain two_chain \ \\ \. finite \ \ finite \ \ (\(k,\) \ \. (\(k',\') \ two_chain_vertical_boundary two_chain. (\a \ {0..1}. \b \ {0..1}. a \ b \ subpath a b \' = \))) \ (common_sudiv_exists (two_chain_horizontal_boundary two_chain) \ \ common_reparam_exists \ (two_chain_horizontal_boundary two_chain)) \ boundary_chain \ \ one_chain = \ \ \ \ (\(k,\)\\. valid_path \)" abbreviation "valid_typeI_division s twoChain \ (\twoCube \ twoChain. typeI_twoCube twoCube) \ gen_division s (cubeImage ` twoChain) \ valid_two_chain twoChain" lemma field_cont_on_typeI_region_cont_on_edges: assumes typeI_twoC: "typeI_twoCube twoC" and field_cont: "continuous_on (cubeImage twoC) F" and member_of_boundary: "(k,\) \ boundary twoC" shows "continuous_on (\ ` {0 .. 1}) F" proof - obtain a b g1 g2 where twoCisTypeI: "a < b" "(\x \ cbox a b. g2 x \ g1 x)" "cubeImage twoC = {(x,y). x \ cbox a b \ y \ cbox (g2 x) (g1 x)}" "twoC = (\(x, y). ((1 - x) * a + x * b, (1 - y) * g2 ((1 - x) * a + x * b) + y * g1 ((1 - x) * a + x * b)))" "g1 piecewise_C1_differentiable_on {a .. b}" "g2 piecewise_C1_differentiable_on {a .. b}" "(\x. twoC(x, 0)) = (\x. (a + (b - a) * x, g2 (a + (b - a) * x)))" "(\y. twoC(1, y)) = (\x. (b, g2 b + x *\<^sub>R (g1 b - g2 b)))" "(\x. twoC(x, 1)) = (\x. (a + (b - a) * x, g1 (a + (b - a) * x)))" "(\y. twoC(0, y)) = (\x. (a, g2 a + x *\<^sub>R (g1 a - g2 a)))" using typeI_twoC and typeI_cube_explicit_spec[of"twoC"] by auto let ?bottom_edge = "(\x. twoC(x, 0))" let ?right_edge = "(\y. twoC(1, y))" let ?top_edge = "(\x. twoC(x, 1))" let ?left_edge = "(\y. twoC(0, y))" let ?Dg1 = "{p. \x. x \ cbox a b \ p = (x, g1(x))}" have line_is_pair_img: "?Dg1 = (\x. (x, g1(x))) ` (cbox a b)" using image_def by auto have field_cont_on_top_edge_image: "continuous_on ?Dg1 F" by (rule continuous_on_subset [OF field_cont]) (auto simp: twoCisTypeI(2) twoCisTypeI(3)) have top_edge_is_compos_of_scal_and_g1: "(\x. twoC(x, 1)) = (\x. (x, g1(x))) \ (\x. a + (b - a) * x)" using twoCisTypeI by auto have Dg1_is_bot_edge_pathimg: "path_image (\x. twoC(x, 1)) = ?Dg1" using line_is_pair_img and top_edge_is_compos_of_scal_and_g1 image_comp path_image_def add_scale_img and twoCisTypeI(1) by (metis (no_types, lifting) cbox_interval) then have cont_on_top: "continuous_on (path_image ?top_edge) F" using field_cont_on_top_edge_image by auto let ?Dg2 = "{p. \x. x \ cbox a b \ p = (x, g2(x))}" have line_is_pair_img: "?Dg2 = (\x. (x, g2(x))) ` (cbox a b)" using image_def by auto have field_cont_on_bot_edge_image: "continuous_on ?Dg2 F" apply (rule continuous_on_subset [OF field_cont]) using twoCisTypeI(2) twoCisTypeI(3) by auto have bot_edge_is_compos_of_scal_and_g2: "(\x. twoC(x, 0)) = (\x. (x, g2(x))) \ (\x. a + (b - a) * x)" using twoCisTypeI by auto have Dg2_is_bot_edge_pathimg: "path_image (\x. twoC(x, 0)) = ?Dg2" using line_is_pair_img and bot_edge_is_compos_of_scal_and_g2 image_comp path_image_def add_scale_img and twoCisTypeI(1) by (metis (no_types, lifting) cbox_interval) then have cont_on_bot: "continuous_on (path_image ?bottom_edge) F" using field_cont_on_bot_edge_image by auto let ?D_left_edge = "{p. \y. y \ cbox (g2 a) (g1 a) \ p = (a, y)}" have field_cont_on_left_edge_image: "continuous_on ?D_left_edge F" apply (rule continuous_on_subset [OF field_cont]) using twoCisTypeI(1) twoCisTypeI(3) by auto have "g2 a \ g1 a" using twoCisTypeI(1) twoCisTypeI(2) by auto then have "(\x. g2 a + (g1 a - g2 a) * x) ` {(0::real)..1} = {g2 a .. g1 a}" using add_scale_img'[of "g2 a" "g1 a"] by blast then have left_eq: "?D_left_edge = ?left_edge ` {0..1}" unfolding twoCisTypeI(10) by(auto simp add: subset_iff image_def set_eq_iff Semiring_Normalization.comm_semiring_1_class.semiring_normalization_rules(7)) then have cont_on_left: "continuous_on (path_image ?left_edge) F" using field_cont_on_left_edge_image path_image_def by (metis left_eq field_cont_on_left_edge_image path_image_def) let ?D_right_edge = "{p. \y. y \ cbox (g2 b) (g1 b) \ p = (b, y)}" have field_cont_on_left_edge_image: "continuous_on ?D_right_edge F" apply (rule continuous_on_subset [OF field_cont]) using twoCisTypeI(1) twoCisTypeI(3) by auto have "g2 b \ g1 b" using twoCisTypeI(1) twoCisTypeI(2) by auto then have "(\x. g2 b + (g1 b - g2 b) * x) ` {(0::real)..1} = {g2 b .. g1 b}" using add_scale_img'[of "g2 b" "g1 b"] by blast then have right_eq: "?D_right_edge = ?right_edge ` {0..1}" unfolding twoCisTypeI(8) by(auto simp add: subset_iff image_def set_eq_iff Semiring_Normalization.comm_semiring_1_class.semiring_normalization_rules(7)) then have cont_on_right: "continuous_on (path_image ?right_edge) F" using field_cont_on_left_edge_image path_image_def by (metis right_eq field_cont_on_left_edge_image path_image_def) have all_edge_cases: "(\ = ?bottom_edge \ \ = ?right_edge \ \ = ?top_edge \ \ = ?left_edge)" using assms by (auto simp add: boundary_def horizontal_boundary_def vertical_boundary_def) show ?thesis apply (simp add: path_image_def[symmetric]) using cont_on_top cont_on_bot cont_on_right cont_on_left all_edge_cases by blast qed lemma typeII_cube_explicit_spec: assumes "typeII_twoCube twoC" shows "\a b g1 g2. a < b \ (\x \ {a .. b}. g2 x \ g1 x) \ cubeImage twoC = {(y, x). x \ {a..b} \ y \ {g2 x .. g1 x}} \ twoC = (\(y, x). ((1 - y) * g2 ((1 - x) * a + x * b) + y * g1 ((1 - x) * a + x * b), (1 - x) * a + x * b)) \ g1 piecewise_C1_differentiable_on {a .. b} \ g2 piecewise_C1_differentiable_on {a .. b} \ (\x. twoC(0, x)) = (\x. (g2 (a + (b - a) * x), a + (b - a) * x)) \ (\y. twoC(y, 1)) = (\x. (g2 b + x *\<^sub>R (g1 b - g2 b), b)) \ (\x. twoC(1, x)) = (\x. (g1 (a + (b - a) * x), a + (b - a) * x)) \ (\y. twoC(y, 0)) = (\x. (g2 a + x *\<^sub>R (g1 a - g2 a), a))" proof - let ?bottom_edge = "(\x. twoC(0, x))" let ?right_edge = "(\y. twoC(y, 1))" let ?top_edge = "(\x. twoC(1, x))" let ?left_edge = "(\y. twoC(y, 0))" obtain a b g1 g2 where twoCisTypeII: "a < b" "(\x \ cbox a b. g2 x \ g1 x)" "cubeImage twoC = {(y, x). x \ cbox a b \ y \ cbox (g2 x) (g1 x)}" "twoC = (\(y, x). ((1 - y) * g2 ((1 - x) * a + x * b) + y * g1 ((1 - x) * a + x * b), (1 - x) * a + x * b))" "g1 piecewise_C1_differentiable_on {a .. b}" "g2 piecewise_C1_differentiable_on {a .. b}" using assms and typeII_twoCubeImg[of"twoC"] by auto have bottom_edge_explicit: "?bottom_edge = (\x. (g2 (a + (b - a) * x), a + (b - a) * x))" by (simp add: twoCisTypeII(4) algebra_simps) have right_edge_explicit: "?right_edge = (\x. (g2 b + x *\<^sub>R (g1 b - g2 b), b))" by (simp add: twoCisTypeII(4) algebra_simps) have top_edge_explicit: "?top_edge = (\x. (g1 (a + (b - a) * x), a + (b - a) * x))" by (simp add: twoCisTypeII(4) algebra_simps) have left_edge_explicit: "?left_edge = (\x. (g2 a + x *\<^sub>R (g1 a - g2 a), a))" by (simp add: twoCisTypeII(4) algebra_simps) show ?thesis using bottom_edge_explicit right_edge_explicit top_edge_explicit left_edge_explicit twoCisTypeII by auto qed lemma typeII_twoCube_smooth_edges: assumes "typeII_twoCube twoC" "(k,\) \ boundary twoC" shows "\ piecewise_C1_differentiable_on {0..1}" proof - let ?bottom_edge = "(\x. twoC(0, x))" let ?right_edge = "(\y. twoC(y, 1))" let ?top_edge = "(\x. twoC(1, x))" let ?left_edge = "(\y. twoC(y, 0))" obtain a b g1 g2 where twoCisTypeII: "a < b" "(\x \ cbox a b. g2 x \ g1 x)" "cubeImage twoC = {(y, x). x \ cbox a b \ y \ cbox (g2 x) (g1 x)}" "twoC = (\(y, x). ((1 - y) * g2 ((1 - x) * a + x * b) + y * g1 ((1 - x) * a + x * b), (1 - x) * a + x * b))" "g1 piecewise_C1_differentiable_on {a .. b}" "g2 piecewise_C1_differentiable_on {a .. b}" "(\x. twoC(0, x)) = (\x. (g2 (a + (b - a) * x), a + (b - a) * x))" "(\y. twoC(y, 1)) = (\x. (g2 b + x *\<^sub>R (g1 b - g2 b), b))" "(\x. twoC(1, x)) = (\x. (g1 (a + (b - a) * x), a + (b - a) * x))" "(\y. twoC(y, 0)) = (\x. (g2 a + x *\<^sub>R (g1 a - g2 a), a))" using assms and typeII_cube_explicit_spec[of"twoC"] by auto have bottom_edge_smooth: "?bottom_edge piecewise_C1_differentiable_on {0..1}" proof - have "\x. (\x. (a + (b - a) * x)) -` {x} = {(x - a)/(b -a)}" using twoCisTypeII(1) by auto then have finite_vimg: "\x. finite({0..1} \ (\x. (a + (b - a) * x))-` {x})" by auto have scale_shif_pw_smth: "(\x. (a + (b - a) * x)) piecewise_C1_differentiable_on {0..1}" using scale_shift_smooth C1_differentiable_imp_piecewise by blast have g2_smooth: "g2 piecewise_C1_differentiable_on (\x. a + (b - a) * x) ` {0..1}" using add_scale_img[OF twoCisTypeII(1)] twoCisTypeII(6) by auto have "(\x. g2 (a + (b - a) * x)) piecewise_C1_differentiable_on {0..1}" - using Cauchy_Integral_Theorem.piecewise_C1_differentiable_compose[OF scale_shif_pw_smth g2_smooth finite_vimg] + using piecewise_C1_differentiable_compose[OF scale_shif_pw_smth g2_smooth finite_vimg] by (auto simp add: o_def) then have "(\x::real. (g2 (a + (b - a) * x), a + (b - a) * x)) piecewise_C1_differentiable_on {0..1}" using all_components_smooth_one_pw_smooth_is_pw_smooth[where f = "(\x::real. (g2 (a + (b - a) * x), a + (b - a) * x))"] by (fastforce simp add: real_pair_basis) then show ?thesis using twoCisTypeII(7) by auto qed have top_edge_smooth: "?top_edge piecewise_C1_differentiable_on {0..1}" proof - have "\x. (\x. (a + (b - a) * x))-` {x} = {(x - a)/(b -a)}" using twoCisTypeII(1) by auto then have finite_vimg: "\x. finite({0..1} \ (\x. (a + (b - a) * x))-` {x})" by auto have scale_shif_pw_smth: "(\x. (a + (b - a) * x)) piecewise_C1_differentiable_on {0..1}" using scale_shift_smooth C1_differentiable_imp_piecewise by blast have g1_smooth: "g1 piecewise_C1_differentiable_on (\x. a + (b - a) * x) ` {0..1}" using add_scale_img[OF twoCisTypeII(1)] twoCisTypeII(5) by auto have "(\x. g1 (a + (b - a) * x)) piecewise_C1_differentiable_on {0..1}" using piecewise_C1_differentiable_compose[OF scale_shif_pw_smth g1_smooth finite_vimg] by (auto simp add: o_def) then have "(\x::real. (g1 (a + (b - a) * x), a + (b - a) * x)) piecewise_C1_differentiable_on {0..1}" using all_components_smooth_one_pw_smooth_is_pw_smooth[where f = "(\x::real. (g1 (a + (b - a) * x), a + (b - a) * x))"] by (fastforce simp add: real_pair_basis) then show ?thesis using twoCisTypeII(9) by auto qed have right_edge_smooth: "?right_edge piecewise_C1_differentiable_on {0..1}" proof - have "(\x. (g2 b + x *\<^sub>R (g1 b - g2 b))) piecewise_C1_differentiable_on {0..1}" by (simp add: C1_differentiable_imp_piecewise) then have "(\x. (g2 b + x *\<^sub>R (g1 b - g2 b), b)) piecewise_C1_differentiable_on {0..1}" using all_components_smooth_one_pw_smooth_is_pw_smooth[of "(1,0)" "(\x. (g2 b + x *\<^sub>R (g1 b - g2 b), b))"] by (auto simp add: real_pair_basis) then show ?thesis using twoCisTypeII(8) by auto qed have left_edge_smooth: "?left_edge piecewise_C1_differentiable_on {0..1}" proof - have 0: "(\x. (g2 a + x *\<^sub>R (g1 a - g2 a))) C1_differentiable_on {0..1}" using C1_differentiable_imp_piecewise by fastforce have "(\x. (g2 a + x *\<^sub>R (g1 a - g2 a), a)) piecewise_C1_differentiable_on {0..1}" using C1_differentiable_imp_piecewise[OF C1_differentiable_on_pair[OF 0 C1_differentiable_on_const[of "a" "{0..1}"]]] by force then show ?thesis using twoCisTypeII(10) by auto qed have "\ = ?bottom_edge \ \ = ?right_edge \ \ = ?top_edge \ \ = ?left_edge" using assms by (auto simp add: boundary_def horizontal_boundary_def vertical_boundary_def) then show ?thesis using left_edge_smooth right_edge_smooth top_edge_smooth bottom_edge_smooth by auto qed lemma field_cont_on_typeII_region_cont_on_edges: assumes typeII_twoC: "typeII_twoCube twoC" and field_cont: "continuous_on (cubeImage twoC) F" and member_of_boundary: "(k,\) \ boundary twoC" shows "continuous_on (\ ` {0 .. 1}) F" proof - obtain a b g1 g2 where twoCisTypeII: "a < b" "(\x \ cbox a b. g2 x \ g1 x)" "cubeImage twoC = {(y, x). x \ cbox a b \ y \ cbox (g2 x) (g1 x)}" "twoC = (\(y, x). ((1 - y) * g2 ((1 - x) * a + x * b) + y * g1 ((1 - x) * a + x * b), (1 - x) * a + x * b))" "g1 piecewise_C1_differentiable_on {a .. b}" "g2 piecewise_C1_differentiable_on {a .. b}" "(\x. twoC(0, x)) = (\x. (g2 (a + (b - a) * x), a + (b - a) * x))" "(\y. twoC(y, 1)) = (\x. (g2 b + x *\<^sub>R (g1 b - g2 b), b))" "(\x. twoC(1, x)) = (\x. (g1 (a + (b - a) * x), a + (b - a) * x))" "(\y. twoC(y, 0)) = (\x. (g2 a + x *\<^sub>R (g1 a - g2 a), a))" using typeII_twoC and typeII_cube_explicit_spec[of"twoC"] by auto let ?bottom_edge = "(\x. twoC(0, x))" let ?right_edge = "(\y. twoC(y, 1))" let ?top_edge = "(\x. twoC(1, x))" let ?left_edge = "(\y. twoC(y, 0))" let ?Dg1 = "{p. \x. x \ cbox a b \ p = (g1(x), x)}" have line_is_pair_img: "?Dg1 = (\x. (g1(x), x)) ` (cbox a b)" using image_def by auto have field_cont_on_top_edge_image: "continuous_on ?Dg1 F" by (rule continuous_on_subset [OF field_cont]) (auto simp: twoCisTypeII(2) twoCisTypeII(3)) have top_edge_is_compos_of_scal_and_g1: "(\x. twoC(1, x)) = (\x. (g1(x), x)) \ (\x. a + (b - a) * x)" using twoCisTypeII by auto have Dg1_is_bot_edge_pathimg: "path_image (\x. twoC(1, x)) = ?Dg1" using line_is_pair_img and top_edge_is_compos_of_scal_and_g1 image_comp path_image_def add_scale_img and twoCisTypeII(1) by (metis (no_types, lifting) cbox_interval) then have cont_on_top: "continuous_on (path_image ?top_edge) F" using field_cont_on_top_edge_image by auto let ?Dg2 = "{p. \x. x \ cbox a b \ p = (g2(x), x)}" have line_is_pair_img: "?Dg2 = (\x. (g2(x), x)) ` (cbox a b)" using image_def by auto have field_cont_on_bot_edge_image: "continuous_on ?Dg2 F" by (rule continuous_on_subset [OF field_cont]) (auto simp add: twoCisTypeII(2) twoCisTypeII(3)) have bot_edge_is_compos_of_scal_and_g2: "(\x. twoC(0, x)) = (\x. (g2(x), x)) \ (\x. a + (b - a) * x)" using twoCisTypeII by auto have Dg2_is_bot_edge_pathimg: "path_image (\x. twoC(0, x)) = ?Dg2" unfolding path_image_def using line_is_pair_img and bot_edge_is_compos_of_scal_and_g2 image_comp add_scale_img [OF \a < b\] by (metis (no_types, lifting) box_real(2)) then have cont_on_bot: "continuous_on (path_image ?bottom_edge) F" using field_cont_on_bot_edge_image by auto let ?D_left_edge = "{p. \y. y \ cbox (g2 a) (g1 a) \ p = (y, a)}" have field_cont_on_left_edge_image: "continuous_on ?D_left_edge F" apply (rule continuous_on_subset [OF field_cont]) using twoCisTypeII(1) twoCisTypeII(3) by auto have "g2 a \ g1 a" using twoCisTypeII(1) twoCisTypeII(2) by auto then have "(\x. g2 a + x * (g1 a - g2 a)) ` {(0::real)..1} = {g2 a .. g1 a}" using add_scale_img'[of "g2 a" "g1 a"] by (auto simp add: ac_simps) with \g2 a \ g1 a\ have left_eq: "?D_left_edge = ?left_edge ` {0..1}" by (simp only: twoCisTypeII(10)) auto then have cont_on_left: "continuous_on (path_image ?left_edge) F" using field_cont_on_left_edge_image path_image_def by (metis left_eq field_cont_on_left_edge_image path_image_def) let ?D_right_edge = "{p. \y. y \ cbox (g2 b) (g1 b) \ p = (y, b)}" have field_cont_on_left_edge_image: "continuous_on ?D_right_edge F" apply (rule continuous_on_subset [OF field_cont]) using twoCisTypeII(1) twoCisTypeII(3) by auto have "g2 b \ g1 b" using twoCisTypeII(1) twoCisTypeII(2) by auto then have "(\x. g2 b + x * (g1 b - g2 b)) ` {(0::real)..1} = {g2 b .. g1 b}" using add_scale_img'[of "g2 b" "g1 b"] by (auto simp add: ac_simps) with \g2 b \ g1 b\ have right_eq: "?D_right_edge = ?right_edge ` {0..1}" by (simp only: twoCisTypeII(8)) auto then have cont_on_right: "continuous_on (path_image ?right_edge) F" using field_cont_on_left_edge_image path_image_def by (metis right_eq field_cont_on_left_edge_image path_image_def) have all_edge_cases: "(\ = ?bottom_edge \ \ = ?right_edge \ \ = ?top_edge \ \ = ?left_edge)" using assms unfolding boundary_def horizontal_boundary_def vertical_boundary_def by blast show ?thesis apply (simp add: path_image_def[symmetric]) using cont_on_top cont_on_bot cont_on_right cont_on_left all_edge_cases by blast qed lemma two_cube_boundary_is_boundary: "boundary_chain (boundary C)" by (auto simp add: boundary_chain_def boundary_def horizontal_boundary_def vertical_boundary_def) lemma common_boundary_subdiv_exists_refl: assumes "\(k,\)\boundary twoC. valid_path \" shows "common_boundary_sudivision_exists (boundary twoC) (boundary twoC)" using assms chain_subdiv_chain_refl common_boundary_sudivision_exists_def two_cube_boundary_is_boundary by blast lemma common_boundary_subdiv_exists_refl': assumes "\(k,\)\C. valid_path \" "boundary_chain (C::(int \ (real \ real \ real)) set)" shows "common_boundary_sudivision_exists (C) (C)" using assms chain_subdiv_chain_refl common_boundary_sudivision_exists_def by blast lemma gen_common_boundary_subdiv_exists_refl_twochain_boundary: assumes "\(k,\)\C. valid_path \" "boundary_chain (C::(int \ (real \ real \ real)) set)" shows "common_sudiv_exists (C) (C)" using assms chain_subdiv_chain_refl common_boundary_sudivision_exists_def common_subdiv_imp_gen_common_subdiv by blast lemma two_chain_boundary_is_boundary_chain: shows "boundary_chain (two_chain_boundary twoChain)" by (simp add: boundary_chain_def two_chain_boundary_def boundary_def horizontal_boundary_def vertical_boundary_def) lemma typeI_edges_are_valid_paths: assumes "typeI_twoCube twoC" "(k,\) \ boundary twoC" shows "valid_path \" using typeI_twoCube_smooth_edges[OF assms] C1_differentiable_imp_piecewise by (auto simp: valid_path_def) lemma typeII_edges_are_valid_paths: assumes "typeII_twoCube twoC" "(k,\) \ boundary twoC" shows "valid_path \" using typeII_twoCube_smooth_edges[OF assms] C1_differentiable_imp_piecewise by (auto simp: valid_path_def) lemma finite_two_chain_vertical_boundary: assumes "finite two_chain" shows "finite (two_chain_vertical_boundary two_chain)" using assms by (simp add: two_chain_vertical_boundary_def vertical_boundary_def) lemma finite_two_chain_horizontal_boundary: assumes "finite two_chain" shows "finite (two_chain_horizontal_boundary two_chain)" using assms by (simp add: two_chain_horizontal_boundary_def horizontal_boundary_def) locale R2 = fixes i j assumes i_is_x_axis: "i = (1::real,0::real)" and j_is_y_axis: "j = (0::real, 1::real)" begin lemma analytically_valid_y: assumes "analytically_valid s F i" shows "(\x. integral UNIV (\y. (partial_vector_derivative F i) (y, x) * (indicator s (y, x)))) \ borel_measurable lborel" proof - have "{(1::real, 0::real), (0, 1)} - {(1, 0)} = {(0::real,1::real)}" by force with assms show ?thesis using assms by(simp add: real_pair_basis analytically_valid_def i_is_x_axis j_is_y_axis) qed lemma analytically_valid_x: assumes "analytically_valid s F j" shows "(\x. integral UNIV (\y. ((partial_vector_derivative F j) (x, y)) * (indicator s (x, y)))) \ borel_measurable lborel" proof - have "{(1::real, 0::real), (0, 1)} - {(0, 1)} = {(1::real, 0::real)}" by force with assms show ?thesis by(simp add: real_pair_basis analytically_valid_def i_is_x_axis j_is_y_axis) qed lemma Greens_thm_type_I: fixes F:: "((real*real) \ (real * real))" and gamma1 gamma2 gamma3 gamma4 :: "(real \ (real * real))" and a:: "real" and b:: "real" and g1:: "(real \ real)" and g2:: "(real \ real)" assumes Dy_def: "Dy_pair = {(x::real,y) . x \ cbox a b \ y \ cbox (g2 x) (g1 x)}" and gamma1_def: "gamma1 = (\x. (a + (b - a) * x, g2(a + (b - a) * x)))" and gamma1_smooth: "gamma1 piecewise_C1_differentiable_on {0..1}" and (*TODO: This should be piecewise smooth*) gamma2_def: "gamma2 = (\x. (b, g2(b) + x *\<^sub>R (g1(b) - g2(b))))" and gamma3_def: "gamma3 = (\x. (a + (b - a) * x, g1(a + (b - a) * x)))" and gamma3_smooth: "gamma3 piecewise_C1_differentiable_on {0..1}" and gamma4_def: "gamma4 = (\x. (a, g2(a) + x *\<^sub>R (g1(a) - g2(a))))" and F_i_analytically_valid: "analytically_valid Dy_pair (\p. F(p) \ i) j" and g2_leq_g1: "\x \ cbox a b. (g2 x) \ (g1 x)" and (*This is needed otherwise what would Dy be?*) a_lt_b: "a < b" shows "(line_integral F {i} gamma1) + (line_integral F {i} gamma2) - (line_integral F {i} gamma3) - (line_integral F {i} gamma4) = (integral Dy_pair (\a. - (partial_vector_derivative (\p. F(p) \ i) j a)))" "line_integral_exists F {i} gamma4" "line_integral_exists F {i} gamma3" "line_integral_exists F {i} gamma2" "line_integral_exists F {i} gamma1" proof - let ?F_b' = "partial_vector_derivative (\a. (F a) \ i) j" have F_first_is_continuous: "continuous_on Dy_pair (\a. F(a) \ i)" using F_i_analytically_valid by (auto simp add: analytically_valid_def) let ?f = "(\x. if x \ (Dy_pair) then (partial_vector_derivative (\a. (F a) \ i) j) x else 0)" have f_lesbegue_integrable: "integrable lborel ?f" using F_i_analytically_valid by (auto simp add: analytically_valid_def indic_ident) have partially_vec_diff: "\a \ Dy_pair. partially_vector_differentiable (\a. (F a) \ i) j a" using F_i_analytically_valid by (auto simp add: analytically_valid_def indicator_def) have x_axis_integral_measurable: "(\x. integral UNIV (\y. ?f(x, y))) \ borel_measurable lborel" proof - have "(\p. (?F_b' p) * indicator (Dy_pair) p) = (\x. if x \ (Dy_pair) then (?F_b') x else 0)" using indic_ident[of "?F_b'"] by auto then have "\x y. ?F_b' (x,y) * indicator (Dy_pair) (x,y) = (\x. if x \ (Dy_pair) then (?F_b') x else 0) (x,y)" by auto then show ?thesis using analytically_valid_x[OF F_i_analytically_valid] by (auto simp add: indicator_def) qed have F_partially_differentiable: "\a\Dy_pair. has_partial_vector_derivative (\x. (F x) \ i) j (?F_b' a) a" using partial_vector_derivative_works partially_vec_diff by fastforce have g1_g2_continuous: "continuous_on (cbox a b) g1" "continuous_on (cbox a b) g2" proof - have shift_scale_cont: "continuous_on {a..b} (\x. (x - a)*(1/(b-a)))" by (intro continuous_intros) have shift_scale_inv: "(\x. a + (b - a) * x) \ (\x. (x - a)*(1/(b-a))) = id" using a_lt_b by (auto simp add: o_def) have img_shift_scale: "(\x. (x - a)*(1/(b-a)))`{a..b} = {0..1}" using a_lt_b apply (auto simp: divide_simps image_iff) apply (rule_tac x="x * (b - a) + a" in bexI) using le_diff_eq by fastforce+ have gamma1_y_component: "(\x. g2(a + (b - a) * x)) = g2 \ (\x.(a + (b - a) * x))" by auto have "continuous_on {0..1} (\x. g2(a + (b - a) * x))" using continuous_on_inner[OF piecewise_C1_differentiable_on_imp_continuous_on[OF gamma1_smooth], of "(\x. j)", OF continuous_on_const] by (simp add: gamma1_def j_is_y_axis) then have "continuous_on {a..b} ((\x. g2(a + (b - a) * x)) \ (\x. (x - a)*(1/(b-a))))" using img_shift_scale continuous_on_compose shift_scale_cont by force then have "continuous_on {a..b} (g2 \ (\x.(a + (b - a) * x)) \ (\x. (x - a)*(1/(b-a))))" using gamma1_y_component by auto then show "continuous_on (cbox a b) g2" using a_lt_b by (simp add: shift_scale_inv) have gamma3_y_component: "(\x. g1(a + (b - a) * x)) = g1 \ (\x.(a + (b - a) * x))" by auto have "continuous_on {0..1} (\x. g1(a + (b - a) * x))" using continuous_on_inner[OF piecewise_C1_differentiable_on_imp_continuous_on[OF gamma3_smooth], of "(\x. j)", OF continuous_on_const] by (simp add: gamma3_def j_is_y_axis) then have "continuous_on {a..b} ((\x. g1(a + (b - a) * x)) \ (\x. (x - a)*(1/(b-a))))" using img_shift_scale continuous_on_compose shift_scale_cont by force then have "continuous_on {a..b} (g1 \ (\x.(a + (b - a) * x)) \ (\x. (x - a)*(1/(b-a))))" using gamma3_y_component by auto then show "continuous_on (cbox a b) g1" using a_lt_b by (simp add: shift_scale_inv) qed have g2_scale_j_contin: "continuous_on (cbox a b) (\x. (0, g2 x))" by (intro continuous_intros g1_g2_continuous) let ?Dg2 = "{p. \x. x \ cbox a b \ p = (x, g2(x))}" have line_is_pair_img: "?Dg2 = (\x. (x, g2(x))) ` (cbox a b)" using image_def by auto have g2_path_continuous: "continuous_on (cbox a b) (\x. (x, g2(x)))" by (intro continuous_intros g1_g2_continuous) have field_cont_on_gamma1_image: "continuous_on ?Dg2 (\a. F(a) \ i)" apply (rule continuous_on_subset [OF F_first_is_continuous]) by (auto simp add: Dy_def g2_leq_g1) have gamma1_is_compos_of_scal_and_g2: "gamma1 = (\x. (x, g2(x))) \ (\x. a + (b - a) * x)" using gamma1_def by auto have add_scale_img: "(\x. a + (b - a) * x) ` {0 .. 1} = {a .. b}" using add_scale_img and a_lt_b by auto then have Dg2_is_gamma1_pathimg: "path_image gamma1 = ?Dg2" by (metis (no_types, lifting) box_real(2) gamma1_is_compos_of_scal_and_g2 image_comp line_is_pair_img path_image_def) have Base_vecs: "i \ Basis" "j \ Basis" "i \ j" using real_pair_basis and i_is_x_axis and j_is_y_axis by auto have gamma1_as_euclid_space_fun: "gamma1 = (\x. (a + (b - a) * x) *\<^sub>R i + (0, g2 (a + (b - a) * x)))" using i_is_x_axis gamma1_def by auto have 0: "line_integral F {i} gamma1 = integral (cbox a b) (\x. F(x, g2(x)) \ i )" "line_integral_exists F {i} gamma1" using line_integral_on_pair_path_strong [OF norm_Basis[OF Base_vecs(1)] _ gamma1_as_euclid_space_fun, of "F"] gamma1_def gamma1_smooth g2_scale_j_contin a_lt_b add_scale_img Dg2_is_gamma1_pathimg and field_cont_on_gamma1_image by (auto simp: pathstart_def pathfinish_def i_is_x_axis) then show "(line_integral_exists F {i} gamma1)" by metis have gamma2_x_const: "\x. gamma2 x \ i = b" by (simp add: i_is_x_axis gamma2_def) have 1: "(line_integral F {i} gamma2) = 0" "(line_integral_exists F {i} gamma2)" using line_integral_on_pair_straight_path[OF gamma2_x_const] straight_path_diffrentiable_x gamma2_def by (auto simp add: mult.commute) then show "(line_integral_exists F {i} gamma2)" by metis have "continuous_on (cbox a b) (\x. F(x, g2(x)) \ i)" using line_is_pair_img and g2_path_continuous and field_cont_on_gamma1_image Topological_Spaces.continuous_on_compose i_is_x_axis j_is_y_axis by auto then have 6: "(\x. F(x, g2(x)) \ i) integrable_on (cbox a b)" using integrable_continuous [of "a" "b" "(\x. F(x, g2(x)) \ i)"] by auto have g1_scale_j_contin: "continuous_on (cbox a b) (\x. (0, g1 x))" by (intro continuous_intros g1_g2_continuous) let ?Dg1 = "{p. \x. x \ cbox a b \ p = (x, g1(x))}" have line_is_pair_img: "?Dg1 = (\x. (x, g1(x)) ) ` (cbox a b)" using image_def by auto have g1_path_continuous: "continuous_on (cbox a b) (\x. (x, g1(x)))" by (intro continuous_intros g1_g2_continuous) have field_cont_on_gamma3_image: "continuous_on ?Dg1 (\a. F(a) \ i)" apply (rule continuous_on_subset [OF F_first_is_continuous]) by (auto simp add: Dy_def g2_leq_g1) have gamma3_is_compos_of_scal_and_g1: "gamma3 = (\x. (x, g1(x))) \ (\x. a + (b - a) * x)" using gamma3_def by auto then have Dg1_is_gamma3_pathimg: "path_image gamma3 = ?Dg1" by (metis (no_types, lifting) box_real(2) image_comp line_is_pair_img local.add_scale_img path_image_def) have Base_vecs: "i \ Basis" "j \ Basis" "i \ j" using real_pair_basis and i_is_x_axis and j_is_y_axis by auto have gamma3_as_euclid_space_fun: "gamma3 = (\x. (a + (b - a) * x) *\<^sub>R i + (0, g1 (a + (b - a) * x)))" using i_is_x_axis gamma3_def by auto have 2: "line_integral F {i} gamma3 = integral (cbox a b) (\x. F(x, g1(x)) \ i )" "line_integral_exists F {i} gamma3" using line_integral_on_pair_path_strong [OF norm_Basis[OF Base_vecs(1)] _ gamma3_as_euclid_space_fun, of "F"] gamma3_def and gamma3_smooth and g1_scale_j_contin and a_lt_b and add_scale_img Dg1_is_gamma3_pathimg and field_cont_on_gamma3_image by (auto simp: pathstart_def pathfinish_def i_is_x_axis) then show "(line_integral_exists F {i} gamma3)" by metis have gamma4_x_const: "\x. gamma4 x \ i = a" using gamma4_def by (auto simp add: real_inner_class.inner_add_left inner_not_same_Basis i_is_x_axis) have 3: "(line_integral F {i} gamma4) = 0" "(line_integral_exists F {i} gamma4)" using line_integral_on_pair_straight_path[OF gamma4_x_const] straight_path_diffrentiable_x gamma4_def by (auto simp add: mult.commute) then show "(line_integral_exists F {i} gamma4)" by metis have "continuous_on (cbox a b) (\x. F(x, g1(x)) \ i)" using line_is_pair_img and g1_path_continuous and field_cont_on_gamma3_image continuous_on_compose i_is_x_axis j_is_y_axis by auto then have 7: "(\x. F(x, g1(x)) \ i) integrable_on (cbox a b)" using integrable_continuous [of "a" "b" "(\x. F(x, g1(x)) \ i)"] by auto have partial_deriv_one_d_integrable: "((\y. ?F_b'(xc, y)) has_integral F(xc,g1(xc)) \ i - F(xc, g2(xc)) \ i) (cbox (g2 xc) (g1 xc))" if "xc \ cbox a b" for xc proof - have "{(xc', y). y \ cbox (g2 xc) (g1 xc) \ xc' = xc} \ Dy_pair" using that by (auto simp add: Dy_def) then show "((\y. ?F_b' (xc, y)) has_integral F(xc, g1 xc) \ i - F(xc, g2 xc) \ i) (cbox (g2 xc) (g1 xc))" using that and Base_vecs and F_partially_differentiable and Dy_def [symmetric] and g2_leq_g1 and fundamental_theorem_of_calculus_partial_vector [of "g2 xc" "g1 xc" "j" "i" "xc *\<^sub>R i" "Dy_pair" "F" "?F_b'" ] by (auto simp add: Groups.ab_semigroup_add_class.add.commute i_is_x_axis j_is_y_axis) qed have partial_deriv_integrable: "(?F_b') integrable_on Dy_pair" by (simp add: F_i_analytically_valid analytically_valid_imp_part_deriv_integrable_on) have 4: "integral Dy_pair ?F_b' = integral (cbox a b) (\x. integral (cbox (g2 x) (g1 x)) (\y. ?F_b' (x, y)))" proof - have x_axis_gauge_integrable: "\x. (\y. ?f(x,y)) integrable_on UNIV" proof - fix x::real have "\x. x \ cbox a b \ (\y. ?f (x, y)) = (\y. 0)" by (auto simp add: Dy_def) then have f_integrable_x_not_in_range: "\x. x \ cbox a b \ (\y. ?f (x, y)) integrable_on UNIV" by (simp add: integrable_0) let ?F_b'_oneD = "(\x. (\y. if y \ (cbox (g2 x) ( g1 x)) then ?F_b' (x,y) else 0))" have f_value_x_in_range: "\x \ cbox a b. ?F_b'_oneD x = (\y. ?f(x,y))" by (auto simp add: Dy_def) have "\x \ cbox a b. ?F_b'_oneD x integrable_on UNIV" using has_integral_integrable integrable_restrict_UNIV partial_deriv_one_d_integrable by blast then have f_integrable_x_in_range: "\x. x \ cbox a b \ (\y. ?f (x, y)) integrable_on UNIV" using f_value_x_in_range by auto show "(\y. ?f (x, y)) integrable_on UNIV" using f_integrable_x_not_in_range and f_integrable_x_in_range by auto qed have arg: "(\a. if a \ Dy_pair then partial_vector_derivative (\a. F a \ i) j a else 0) = (\x. if x \ Dy_pair then if x \ Dy_pair then partial_vector_derivative (\a. F a \ i) j x else 0 else 0)" by auto have arg2: "Dy_pair = {(x, y). (\i\Basis. a \ i \ x \ i \ x \ i \ b \ i) \ (\i\Basis. g2 x \ i \ y \ i \ y \ i \ g1 x \ i)}" using Dy_def by auto have arg3: "\ x. x \ Dy_pair \ (\x. if x \ Dy_pair then partial_vector_derivative (\a. F a \ i) j x else 0) x = (\x. partial_vector_derivative (\a. F a \ i) j x) x" by auto have arg4: "\x. x \ (cbox a b) \ (\x. integral (cbox (g2 x) (g1 x)) (\y. if (x, y) \ Dy_pair then partial_vector_derivative (\a. F a \ i) j (x, y) else 0)) x = (\x. integral (cbox (g2 x) (g1 x)) (\y. partial_vector_derivative (\a. F a \ i) j (x, y))) x" apply (simp add: Dy_def) by (smt Henstock_Kurzweil_Integration.integral_cong atLeastAtMost_iff) show ?thesis using gauge_integral_Fubini_curve_bounded_region_x [OF f_lesbegue_integrable x_axis_gauge_integrable x_axis_integral_measurable arg arg2] Henstock_Kurzweil_Integration.integral_cong[OF arg3, of "Dy_pair" "(\x. x)"] Henstock_Kurzweil_Integration.integral_cong[OF arg4, of "cbox a b" "(\x. x)"] by auto qed have 5: "(integral Dy_pair (\a. (?F_b' a)) = integral (cbox a b) (\x. F(x, g1(x)) \ i - F(x, g2(x)) \ i))" using 4 Henstock_Kurzweil_Integration.integral_cong partial_deriv_one_d_integrable integrable_def by (smt integral_unique) show "(line_integral F {i} gamma1) + (line_integral F {i} gamma2) - (line_integral F {i} gamma3) - (line_integral F {i} gamma4) = (integral Dy_pair (\a. - (?F_b' a)))" using "0"(1) "1"(1) "2"(1) "3"(1) 5 "6" "7" by (simp add: Henstock_Kurzweil_Integration.integral_diff) qed theorem Greens_thm_type_II: fixes F:: "((real*real) \ (real * real))" and gamma4 gamma3 gamma2 gamma1 :: "(real \ (real * real))" and a:: "real" and b:: "real" and g1:: "(real \ real)" and g2:: "(real \ real)" assumes Dx_def: "Dx_pair = {(x::real,y) . y \ cbox a b \ x \ cbox (g2 y) (g1 y)}" and gamma4_def: "gamma4 = (\x. (g2(a + (b - a) * x), a + (b - a) * x))" and gamma4_smooth: "gamma4 piecewise_C1_differentiable_on {0..1}" and (*TODO: This should be piecewise smooth*) gamma3_def: "gamma3 = (\x. (g2(b) + x *\<^sub>R (g1(b) - g2(b)), b))" and gamma2_def: "gamma2 = (\x. (g1(a + (b - a) * x), a + (b - a) * x))" and gamma2_smooth: "gamma2 piecewise_C1_differentiable_on {0..1}" and gamma1_def: "gamma1 = (\x. (g2(a) + x *\<^sub>R (g1(a) - g2(a)), a))" and F_j_analytically_valid: "analytically_valid Dx_pair (\p. F(p) \ j) i" and g2_leq_g1: "\x \ cbox a b. (g2 x) \ (g1 x)" and (*This is needed otherwise what would Dy be?*) a_lt_b: "a < b" shows "-(line_integral F {j} gamma4) - (line_integral F {j} gamma3) + (line_integral F {j} gamma2) + (line_integral F {j} gamma1) = (integral Dx_pair (\a. (partial_vector_derivative (\a. (F a) \ j) i a)))" "line_integral_exists F {j} gamma4" "line_integral_exists F {j} gamma3" "line_integral_exists F {j} gamma2" "line_integral_exists F {j} gamma1" proof - let ?F_a' = "partial_vector_derivative (\a. (F a) \ j) i" have F_first_is_continuous: "continuous_on Dx_pair (\a. F(a) \ j)" using F_j_analytically_valid by (auto simp add: analytically_valid_def) let ?f = "(\x. if x \ (Dx_pair) then (partial_vector_derivative (\a. (F a) \ j) i) x else 0)" have f_lesbegue_integrable: "integrable lborel ?f" using F_j_analytically_valid by (auto simp add: analytically_valid_def indic_ident) have partially_vec_diff: "\a \ Dx_pair. partially_vector_differentiable (\a. (F a) \ j) i a" using F_j_analytically_valid by (auto simp add: analytically_valid_def indicator_def) have "\x y. ?F_a' (x,y) * indicator (Dx_pair) (x,y) = (\x. if x \ (Dx_pair) then ?F_a' x else 0) (x,y)" using indic_ident[of "?F_a'"] by auto then have y_axis_integral_measurable: "(\x. integral UNIV (\y. ?f(y, x))) \ borel_measurable lborel" using analytically_valid_y[OF F_j_analytically_valid] by (auto simp add: indicator_def) have F_partially_differentiable: "\a\Dx_pair. has_partial_vector_derivative (\x. (F x) \ j) i (?F_a' a) a" using partial_vector_derivative_works partially_vec_diff by fastforce have g1_g2_continuous: "continuous_on (cbox a b) g1" "continuous_on (cbox a b) g2" proof - have shift_scale_cont: "continuous_on {a..b} (\x. (x - a)*(1/(b-a)))" by (intro continuous_intros) have shift_scale_inv: "(\x. a + (b - a) * x) \ (\x. (x - a)*(1/(b-a))) = id" using a_lt_b by (auto simp add: o_def) have img_shift_scale: "(\x. (x - a)*(1/(b-a)))`{a..b} = {0..1}" apply (auto simp: divide_simps image_iff) apply (rule_tac x="x * (b - a) + a" in bexI) using a_lt_b by (auto simp: algebra_simps mult_le_cancel_left affine_ineq) have "continuous_on {0..1} (\x. g2(a + (b - a) * x))" using continuous_on_inner[OF piecewise_C1_differentiable_on_imp_continuous_on[OF gamma4_smooth], of "(\x. i)", OF continuous_on_const] by (simp add: gamma4_def i_is_x_axis) then have "continuous_on {a..b} ((\x. g2(a + (b - a) * x)) \ (\x. (x - a)*(1/(b-a))))" using img_shift_scale continuous_on_compose shift_scale_cont by force then show "continuous_on (cbox a b) g2" using a_lt_b by (simp add: shift_scale_inv) have "continuous_on {0..1} (\x. g1(a + (b - a) * x))" using continuous_on_inner[OF piecewise_C1_differentiable_on_imp_continuous_on[OF gamma2_smooth], of "(\x. i)", OF continuous_on_const] by (simp add: gamma2_def i_is_x_axis) then have "continuous_on {a..b} ((\x. g1(a + (b - a) * x)) \ (\x. (x - a)*(1/(b-a))))" using img_shift_scale continuous_on_compose shift_scale_cont by force then show "continuous_on (cbox a b) g1" using a_lt_b by (simp add: shift_scale_inv) qed have g2_scale_i_contin: "continuous_on (cbox a b) (\x. (g2 x, 0))" by (intro continuous_intros g1_g2_continuous) let ?Dg2 = "{p. \x. x \ cbox a b \ p = (g2(x), x)}" have line_is_pair_img: "?Dg2 = (\x. (g2(x), x) ) ` (cbox a b)" using image_def by auto have g2_path_continuous: "continuous_on (cbox a b) (\x. (g2(x), x))" by (intro continuous_intros g1_g2_continuous) have field_cont_on_gamma4_image: "continuous_on ?Dg2 (\a. F(a) \ j)" by (rule continuous_on_subset [OF F_first_is_continuous]) (auto simp: Dx_def g2_leq_g1) have gamma4_is_compos_of_scal_and_g2: "gamma4 = (\x. (g2(x), x)) \ (\x. a + (b - a) * x)" using gamma4_def by auto have add_scale_img: "(\x. a + (b - a) * x) ` {0 .. 1} = {a .. b}" using add_scale_img and a_lt_b by auto then have Dg2_is_gamma4_pathimg: "path_image gamma4 = ?Dg2" using line_is_pair_img and gamma4_is_compos_of_scal_and_g2 image_comp path_image_def by (metis (no_types, lifting) cbox_interval) have Base_vecs: "i \ Basis" "j \ Basis" "i \ j" using real_pair_basis and i_is_x_axis and j_is_y_axis by auto have gamma4_as_euclid_space_fun: "gamma4 = (\x. (a + (b - a) * x) *\<^sub>R j + (g2 (a + (b - a) * x), 0))" using j_is_y_axis gamma4_def by auto have 0: "(line_integral F {j} gamma4) = integral (cbox a b) (\x. F(g2(x), x) \ j)" "line_integral_exists F {j} gamma4" using line_integral_on_pair_path_strong [OF norm_Basis[OF Base_vecs(2)] _ gamma4_as_euclid_space_fun] gamma4_def gamma4_smooth g2_scale_i_contin a_lt_b add_scale_img Dg2_is_gamma4_pathimg and field_cont_on_gamma4_image by (auto simp: pathstart_def pathfinish_def j_is_y_axis) then show "line_integral_exists F {j} gamma4" by metis have gamma3_y_const: "\x. gamma3 x \ j = b" by (simp add: gamma3_def j_is_y_axis) have 1: "(line_integral F {j} gamma3) = 0" "(line_integral_exists F {j} gamma3)" using line_integral_on_pair_straight_path[OF gamma3_y_const] straight_path_diffrentiable_y gamma3_def by (auto simp add: mult.commute) then show "line_integral_exists F {j} gamma3" by auto have "continuous_on (cbox a b) (\x. F(g2(x), x) \ j)" by (smt Collect_mono_iff continuous_on_compose2 continuous_on_eq field_cont_on_gamma4_image g2_path_continuous line_is_pair_img) then have 6: "(\x. F(g2(x), x) \ j) integrable_on (cbox a b)" using integrable_continuous by blast have g1_scale_i_contin: "continuous_on (cbox a b) (\x. (g1 x, 0))" by (intro continuous_intros g1_g2_continuous) let ?Dg1 = "{p. \x. x \ cbox a b \ p = (g1(x), x)}" have line_is_pair_img: "?Dg1 = (\x. (g1(x), x) ) ` (cbox a b)" using image_def by auto have g1_path_continuous: "continuous_on (cbox a b) (\x. (g1(x), x))" by (intro continuous_intros g1_g2_continuous) have field_cont_on_gamma2_image: "continuous_on ?Dg1 (\a. F(a) \ j)" by (rule continuous_on_subset [OF F_first_is_continuous]) (auto simp: Dx_def g2_leq_g1) have "gamma2 = (\x. (g1(x), x)) \ (\x. a + (b - a) * x)" using gamma2_def by auto then have Dg1_is_gamma2_pathimg: "path_image gamma2 = ?Dg1" using line_is_pair_img image_comp path_image_def add_scale_img by (metis (no_types, lifting) cbox_interval) have Base_vecs: "i \ Basis" "j \ Basis" "i \ j" using real_pair_basis and i_is_x_axis and j_is_y_axis by auto have gamma2_as_euclid_space_fun: "gamma2 = (\x. (a + (b - a) * x) *\<^sub>R j + (g1 (a + (b - a) * x), 0))" using j_is_y_axis gamma2_def by auto have 2: "(line_integral F {j} gamma2) = integral (cbox a b) (\x. F(g1(x), x) \ j)" "(line_integral_exists F {j} gamma2)" using line_integral_on_pair_path_strong [OF norm_Basis[OF Base_vecs(2)] _ gamma2_as_euclid_space_fun] gamma2_def and gamma2_smooth and g1_scale_i_contin and a_lt_b and add_scale_img Dg1_is_gamma2_pathimg and field_cont_on_gamma2_image by (auto simp: pathstart_def pathfinish_def j_is_y_axis) then show "line_integral_exists F {j} gamma2" by metis have gamma1_y_const: "\x. gamma1 x \ j = a" using gamma1_def by (auto simp add: real_inner_class.inner_add_left euclidean_space_class.inner_not_same_Basis j_is_y_axis) have 3: "(line_integral F {j} gamma1) = 0" "(line_integral_exists F {j} gamma1)" using line_integral_on_pair_straight_path[OF gamma1_y_const] straight_path_diffrentiable_y gamma1_def by (auto simp add: mult.commute) then show "line_integral_exists F {j} gamma1" by auto have "continuous_on (cbox a b) (\x. F(g1(x), x) \ j)" by (smt Collect_mono_iff continuous_on_compose2 continuous_on_eq field_cont_on_gamma2_image g1_path_continuous line_is_pair_img) then have 7: "(\x. F(g1(x), x) \ j) integrable_on (cbox a b)" using integrable_continuous [of "a" "b" "(\x. F(g1(x), x) \ j)"] by auto have partial_deriv_one_d_integrable: "((\y. ?F_a'(y, xc)) has_integral F(g1(xc), xc) \ j - F(g2(xc), xc) \ j) (cbox (g2 xc) (g1 xc))" if "xc \ cbox a b" for xc::real proof - have "{(y, xc'). y \ cbox (g2 xc) (g1 xc) \ xc' = xc} \ Dx_pair" using that by (auto simp add: Dx_def) then show ?thesis using that and Base_vecs and F_partially_differentiable and Dx_def [symmetric] and g2_leq_g1 and fundamental_theorem_of_calculus_partial_vector [of "g2 xc" "g1 xc" "i" "j" "xc *\<^sub>R j" "Dx_pair" "F" "?F_a'" ] by (auto simp add: Groups.ab_semigroup_add_class.add.commute i_is_x_axis j_is_y_axis) qed have "?f integrable_on UNIV" by (simp add: f_lesbegue_integrable integrable_on_lborel) then have partial_deriv_integrable: "?F_a' integrable_on Dx_pair" using integrable_restrict_UNIV by auto have 4: "integral Dx_pair ?F_a' = integral (cbox a b) (\x. integral (cbox (g2 x) (g1 x)) (\y. ?F_a' (y, x)))" proof - have y_axis_gauge_integrable: "(\y. ?f(y, x)) integrable_on UNIV" for x proof - let ?F_a'_oneD = "(\x. (\y. if y \ (cbox (g2 x) ( g1 x)) then ?F_a' (y, x) else 0))" have "\x. x \ cbox a b \ (\y. ?f (y, x)) = (\y. 0)" by (auto simp add: Dx_def) then have f_integrable_x_not_in_range: "\x. x \ cbox a b \ (\y. ?f (y, x)) integrable_on UNIV" by (simp add: integrable_0) have "\x \ cbox a b. ?F_a'_oneD x = (\y. ?f(y, x))" using g2_leq_g1 by (auto simp add: Dx_def) moreover have "\x \ cbox a b. ?F_a'_oneD x integrable_on UNIV" using has_integral_integrable integrable_restrict_UNIV partial_deriv_one_d_integrable by blast ultimately have "\x. x \ cbox a b \ (\y. ?f (y, x)) integrable_on UNIV" by auto then show "(\y. ?f (y, x)) integrable_on UNIV" using f_integrable_x_not_in_range by auto qed have arg: "(\a. if a \ Dx_pair then partial_vector_derivative (\a. F a \ j) i a else 0) = (\x. if x \ Dx_pair then if x \ Dx_pair then partial_vector_derivative (\a. F a \ j) i x else 0 else 0)" by auto have arg2: "Dx_pair = {(y, x). (\i\Basis. a \ i \ x \ i \ x \ i \ b \ i) \ (\i\Basis. g2 x \ i \ y \ i \ y \ i \ g1 x \ i)}" using Dx_def by auto have arg3: "\ x. x \ Dx_pair \ (\x. if x \ Dx_pair then partial_vector_derivative (\a. F a \ j) i x else 0) x = (\x. partial_vector_derivative (\a. F a \ j) i x) x" by auto have arg4: "\x. x \ (cbox a b) \ (\x. integral (cbox (g2 x) (g1 x)) (\y. if (y, x) \ Dx_pair then partial_vector_derivative (\a. F a \ j) i (y, x) else 0)) x = (\x. integral (cbox (g2 x) (g1 x)) (\y. partial_vector_derivative (\a. F a \ j) i (y, x))) x" apply (clarsimp simp: Dx_def) by (smt Henstock_Kurzweil_Integration.integral_cong atLeastAtMost_iff) show ?thesis using gauge_integral_Fubini_curve_bounded_region_y [OF f_lesbegue_integrable y_axis_gauge_integrable y_axis_integral_measurable arg arg2] Henstock_Kurzweil_Integration.integral_cong[OF arg3, of "Dx_pair" "(\x. x)"] Henstock_Kurzweil_Integration.integral_cong[OF arg4, of "cbox a b" "(\x. x)"] by auto qed have "((integral Dx_pair (\a. (?F_a' a))) = integral (cbox a b) (\x. F(g1(x), x) \ j - F(g2(x), x) \ j))" using 4 Henstock_Kurzweil_Integration.integral_cong partial_deriv_one_d_integrable integrable_def by (smt integral_unique) then have "integral Dx_pair (\a. - (?F_a' a)) = - integral (cbox a b) (\x. F(g1(x), x) \ j - F(g2(x), x) \ j)" using partial_deriv_integrable and integral_neg by auto then have 5: "integral Dx_pair (\a. - (?F_a' a)) = integral (cbox a b) (\x. ( F(g2(x), x) \ j - F(g1(x), x) \ j))" using 6 7 and integral_neg [of _ "(\x. F(g1 x, x) \ j - F(g2 x, x) \ j)"] by auto have "(line_integral F {j} gamma4) + (line_integral F {j} gamma3) - (line_integral F {j} gamma2) - (line_integral F {j} gamma1) = (integral Dx_pair (\a. -(?F_a' a)))" using 0 1 2 3 5 6 7 Henstock_Kurzweil_Integration.integral_diff[of "(\x. F(g2(x), x) \ j)" "cbox a b" "(\x. F(g1(x), x) \ j)"] by (auto) then show "-(line_integral F {j} gamma4) - (line_integral F {j} gamma3) + (line_integral F {j} gamma2) + (line_integral F {j} gamma1) = (integral Dx_pair (\a. (?F_a' a)))" by (simp add: \integral Dx_pair (\a. - ?F_a' a) = - integral (cbox a b) (\x. F(g1 x, x) \ j - F(g2 x, x) \ j)\ \integral Dx_pair ?F_a' = integral (cbox a b) (\x. F(g1 x, x) \ j - F(g2 x, x) \ j)\) qed end locale green_typeII_cube = R2 + fixes twoC F assumes two_cube: "typeII_twoCube twoC" and valid_two_cube: "valid_two_cube twoC" and f_analytically_valid: "analytically_valid (cubeImage twoC) (\x. (F x) \ j) i" begin lemma GreenThm_typeII_twoCube: shows "integral (cubeImage twoC) (\a. partial_vector_derivative (\x. (F x) \ j) i a) = one_chain_line_integral F {j} (boundary twoC)" "\(k,\) \ boundary twoC. line_integral_exists F {j} \" proof - let ?bottom_edge = "(\x. twoC(x, 0))" let ?right_edge = "(\y. twoC(1, y))" let ?top_edge = "(\x. twoC(x, 1))" let ?left_edge = "(\y. twoC(0, y))" have line_integral_around_boundary: "one_chain_line_integral F {j} (boundary twoC) = line_integral F {j} ?bottom_edge + line_integral F {j} ?right_edge - line_integral F {j} ?top_edge - line_integral F {j} ?left_edge" proof (simp add: one_chain_line_integral_def horizontal_boundary_def vertical_boundary_def boundary_def) have finite1: "finite {(- 1::int, \y. twoC (0, y)), (1, \y. twoC (1, y)), (- 1, \x. twoC (x, 1))}" by auto then have sum_step1: "(\(k, g)\{(- (1::int), \y. twoC (0, y)), (1, \y. twoC (1, y)), (1, \x. twoC (x, 0)), (- 1, \x. twoC (x, 1))}. k * line_integral F {j} g) = line_integral F {j} (\x. twoC (x, 0)) + (\(k, g)\{(- (1::int), \y. twoC (0, y)), (1, \y. twoC (1, y)), (- 1, \x. twoC (x, 1))}. k * line_integral F {j} g)" using sum.insert_remove [OF finite1] using valid_two_cube apply (simp only: one_chain_line_integral_def horizontal_boundary_def vertical_boundary_def boundary_def valid_two_cube_def) by (auto simp: card_insert_if split: if_split_asm) have three_distinct_edges: "card {(- 1::int, \y. twoC (0, y)), (1, \y. twoC (1, y)), (- 1, \x. twoC (x, 1))} = 3" using valid_two_cube apply (simp add: one_chain_line_integral_def horizontal_boundary_def vertical_boundary_def boundary_def valid_two_cube_def) by (auto simp: card_insert_if split: if_split_asm) have finite2: "finite {(- 1::int, \y. twoC (0, y)), (1, \y. twoC (1, y))}" by auto then have sum_step2: "(\(k, g)\{(- (1::int), \y. twoC (0, y)), (1, \y. twoC (1, y)), (-1, \x. twoC (x, 1))}. k * line_integral F {j} g) = (- line_integral F {j} (\x. twoC (x, 1))) + (\(k, g)\{(- (1::int), \y. twoC (0, y)), (1, \y. twoC (1, y))}. k * line_integral F {j} g)" using sum.insert_remove [OF finite2] three_distinct_edges by (auto simp: card_insert_if split: if_split_asm) have two_distinct_edges: "card {(- 1::int, \y. twoC (0, y)), (1, \y. twoC (1, y))} = 2" using three_distinct_edges by (simp add: one_chain_line_integral_def horizontal_boundary_def vertical_boundary_def boundary_def) have finite3: "finite {(- 1::int, \y. twoC (0, y))}" by auto then have sum_step3: "(\(k, g)\{(- (1::int), \y. twoC (0, y)), (1, \y. twoC (1, y))}. k * line_integral F {j} g) = ( line_integral F {j} (\y. twoC (1, y))) + (\(k, g)\{(- (1::real), \y. twoC (0, y))}. k * line_integral F {j} g)" using sum.insert_remove [OF finite2] three_distinct_edges by (auto simp: card_insert_if split: if_split_asm) show "(\x\{(- 1::int, \y. twoC (0, y)), (1, \y. twoC (1, y)), (1, \x. twoC (x, 0)), (- 1, \x. twoC (x, 1))}. case x of (k, g) \ k * line_integral F {j} g) = line_integral F {j} (\x. twoC (x, 0)) + line_integral F {j} (\y. twoC (1, y)) - line_integral F {j} (\x. twoC (x, 1)) - line_integral F {j} (\y. twoC (0, y))" using sum_step1 sum_step2 sum_step3 by auto qed obtain a b g1 g2 where twoCisTypeII: "a < b" "(\x \ cbox a b. g2 x \ g1 x)" "cubeImage twoC = {(y, x). x \ cbox a b \ y \ cbox (g2 x) (g1 x)}" "twoC = (\(y, x). ((1 - y) * g2 ((1 - x) * a + x * b) + y * g1 ((1 - x) * a + x * b), (1 - x) * a + x * b))" "g1 piecewise_C1_differentiable_on {a .. b}" "g2 piecewise_C1_differentiable_on {a .. b}" using two_cube and typeII_twoCubeImg[of"twoC"] by auto have left_edge_explicit: "?left_edge = (\x. (g2 (a + (b - a) * x), a + (b - a) * x))" by (simp add: twoCisTypeII(4) algebra_simps) have left_edge_smooth: "?left_edge piecewise_C1_differentiable_on {0..1}" proof - have "\x. (\x. (a + (b - a) * x))-` {x} = {(x - a)/(b -a)}" using twoCisTypeII(1) by(auto simp add: Set.vimage_def) then have finite_vimg: "\x. finite({0..1} \ (\x. (a + (b - a) * x))-` {x})" by auto have scale_shif_smth: "(\x. (a + (b - a) * x)) C1_differentiable_on {0..1}" using scale_shift_smooth by auto then have scale_shif_pw_smth: "(\x. (a + (b - a) * x)) piecewise_C1_differentiable_on {0..1}" using C1_differentiable_imp_piecewise by blast have g2_smooth: "g2 piecewise_C1_differentiable_on (\x. a + (b - a) * x) ` {0..1}" using add_scale_img[OF twoCisTypeII(1)] twoCisTypeII(6) by auto have "(\x. g2 (a + (b - a) * x)) piecewise_C1_differentiable_on {0..1}" - using Cauchy_Integral_Theorem.piecewise_C1_differentiable_compose[OF scale_shif_pw_smth g2_smooth finite_vimg] + using piecewise_C1_differentiable_compose[OF scale_shif_pw_smth g2_smooth finite_vimg] by (auto simp add: o_def) then have "(\x::real. (g2 (a + (b - a) * x), a + (b - a) * x)) piecewise_C1_differentiable_on {0..1}" using all_components_smooth_one_pw_smooth_is_pw_smooth[where f = "(\x::real. (g2 (a + (b - a) * x), a + (b - a) * x))"] by (fastforce simp add: real_pair_basis) then show ?thesis using left_edge_explicit by auto qed have top_edge_explicit: "?top_edge = (\x. (g2 b + x *\<^sub>R (g1 b - g2 b), b))" and right_edge_explicit: "?right_edge = (\x. (g1 (a + (b - a) * x), a + (b - a) * x))" by (simp_all add: twoCisTypeII(4) algebra_simps) have right_edge_smooth: "?right_edge piecewise_C1_differentiable_on {0..1}" proof - have "\x. (\x. (a + (b - a) * x))-` {x} = {(x - a)/(b -a)}" using twoCisTypeII(1) by(auto simp add: Set.vimage_def) then have finite_vimg: "\x. finite({0..1} \ (\x. (a + (b - a) * x))-` {x})" by auto then have scale_shif_pw_smth: "(\x. (a + (b - a) * x)) piecewise_C1_differentiable_on {0..1}" using C1_differentiable_imp_piecewise [OF scale_shift_smooth] by auto have g1_smooth: "g1 piecewise_C1_differentiable_on (\x. a + (b - a) * x) ` {0..1}" using add_scale_img[OF twoCisTypeII(1)] twoCisTypeII(5) by auto have "(\x. g1 (a + (b - a) * x)) piecewise_C1_differentiable_on {0..1}" using piecewise_C1_differentiable_compose[OF scale_shif_pw_smth g1_smooth finite_vimg] by (auto simp add: o_def) then have "(\x::real. (g1 (a + (b - a) * x), a + (b - a) * x)) piecewise_C1_differentiable_on {0..1}" using all_components_smooth_one_pw_smooth_is_pw_smooth[where f = "(\x::real. (g1 (a + (b - a) * x), a + (b - a) * x))"] by (fastforce simp add: real_pair_basis) then show ?thesis using right_edge_explicit by auto qed have bottom_edge_explicit: "?bottom_edge = (\x. (g2 a + x *\<^sub>R (g1 a - g2 a), a))" by (simp add: twoCisTypeII(4) algebra_simps) show "integral (cubeImage twoC) (\a. partial_vector_derivative (\x. (F x) \ j) i a) = one_chain_line_integral F {j} (boundary twoC)" using Greens_thm_type_II[OF twoCisTypeII(3) left_edge_explicit left_edge_smooth top_edge_explicit right_edge_explicit right_edge_smooth bottom_edge_explicit f_analytically_valid twoCisTypeII(2) twoCisTypeII(1)] line_integral_around_boundary by auto have "line_integral_exists F {j} \" if "(k,\) \ boundary twoC" for k \ proof - have "line_integral_exists F {j} (\y. twoC (0, y))" "line_integral_exists F {j} (\x. twoC (x, 1))" "line_integral_exists F {j} (\y. twoC (1, y))" "line_integral_exists F {j} (\x. twoC (x, 0))" using Greens_thm_type_II[OF twoCisTypeII(3) left_edge_explicit left_edge_smooth top_edge_explicit right_edge_explicit right_edge_smooth bottom_edge_explicit f_analytically_valid twoCisTypeII(2) twoCisTypeII(1)] line_integral_around_boundary by auto then show "line_integral_exists F {j} \" using that by (auto simp add: boundary_def horizontal_boundary_def vertical_boundary_def) qed then show "\(k,\) \ boundary twoC. line_integral_exists F {j} \" by auto qed lemma line_integral_exists_on_typeII_Cube_boundaries': assumes "(k,\) \ boundary twoC" shows "line_integral_exists F {j} \" using assms GreenThm_typeII_twoCube(2) by blast end locale green_typeII_chain = R2 + fixes F two_chain s assumes valid_typeII_div: "valid_typeII_division s two_chain" and F_anal_valid: "\twoC \ two_chain. analytically_valid (cubeImage twoC) (\x. (F x) \ j) i" begin lemma two_chain_valid_valid_cubes: "\two_cube \ two_chain. valid_two_cube two_cube" using valid_typeII_div by (auto simp add: valid_two_chain_def) lemma typeII_chain_line_integral_exists_boundary': shows "\(k,\) \ two_chain_vertical_boundary two_chain. line_integral_exists F {j} \" proof - have integ_exis: "\(k,\) \ two_chain_boundary two_chain. line_integral_exists F {j} \" using green_typeII_cube.line_integral_exists_on_typeII_Cube_boundaries'[of i j] F_anal_valid valid_typeII_div apply(auto simp add: two_chain_boundary_def) using R2_axioms green_typeII_cube_axioms_def green_typeII_cube_def two_chain_valid_valid_cubes by blast then show integ_exis_vert: "\(k,\) \ two_chain_vertical_boundary two_chain. line_integral_exists F {j} \" by (simp add: two_chain_boundary_def two_chain_vertical_boundary_def boundary_def) qed lemma typeII_chain_line_integral_exists_boundary'': "\(k,\) \ two_chain_horizontal_boundary two_chain. line_integral_exists F {j} \" proof - have integ_exis: "\(k,\) \ two_chain_boundary two_chain. line_integral_exists F {j} \" using green_typeII_cube.line_integral_exists_on_typeII_Cube_boundaries'[of i j] valid_typeII_div apply (simp add: two_chain_boundary_def boundary_def) using F_anal_valid R2_axioms green_typeII_cube_axioms_def green_typeII_cube_def two_chain_valid_valid_cubes by fastforce then show integ_exis_vert: "\(k,\) \ two_chain_horizontal_boundary two_chain. line_integral_exists F {j} \" by (simp add: two_chain_boundary_def two_chain_horizontal_boundary_def boundary_def) qed lemma typeII_cube_line_integral_exists_boundary: "\(k,\) \ two_chain_boundary two_chain. line_integral_exists F {j} \" using valid_typeII_div typeII_chain_line_integral_exists_boundary' typeII_chain_line_integral_exists_boundary'' apply (auto simp add: two_chain_boundary_def two_chain_horizontal_boundary_def two_chain_vertical_boundary_def) using boundary_def by auto lemma type_II_chain_horiz_bound_valid: "\(k,\) \ two_chain_horizontal_boundary two_chain. valid_path \" using valid_typeII_div typeII_edges_are_valid_paths by (force simp add: two_chain_boundary_def two_chain_horizontal_boundary_def boundary_def) lemma type_II_chain_vert_bound_valid: (*This and the previous one need to be used in all proofs*) "\(k,\) \ two_chain_vertical_boundary two_chain. valid_path \" using typeII_edges_are_valid_paths valid_typeII_div by (force simp add: two_chain_boundary_def two_chain_vertical_boundary_def boundary_def) lemma members_of_only_horiz_div_line_integrable': assumes "only_horizontal_division one_chain two_chain" "(k::int, \)\one_chain" "(k::int, \)\one_chain" "finite two_chain" "\two_cube \ two_chain. valid_two_cube two_cube" shows "line_integral_exists F {j} \" proof - have integ_exis: "\(k,\) \ two_chain_boundary two_chain. line_integral_exists F {j} \" using typeII_cube_line_integral_exists_boundary by blast have integ_exis_vert: "\(k,\) \ two_chain_vertical_boundary two_chain. line_integral_exists F {j} \" using typeII_chain_line_integral_exists_boundary' assms by auto have integ_exis_horiz: "(\k \. (\(k', \') \ two_chain_horizontal_boundary two_chain. \a\{0..1}. \b\{0..1}. a \ b \ subpath a b \' = \) \ line_integral_exists F {j} \)" using integ_exis type_II_chain_horiz_bound_valid using line_integral_exists_subpath[of "F" "{j}"] by (fastforce simp add: two_chain_boundary_def two_chain_horizontal_boundary_def two_chain_vertical_boundary_def boundary_def) obtain \ \ where hv_props: "finite \" "(\(k,\) \ \. (\(k', \') \ two_chain_horizontal_boundary two_chain. (\a \ {0 .. 1}. \b \ {0..1}. a \ b \ subpath a b \' = \)))" "one_chain = \ \ \" "((common_sudiv_exists (two_chain_vertical_boundary two_chain) \) \ common_reparam_exists \ (two_chain_vertical_boundary two_chain))" "finite \" "boundary_chain \" "\(k,\)\\. valid_path \" using assms(1) unfolding only_horizontal_division_def by blast have finite_j: "finite {j}" by auto show "line_integral_exists F {j} \" proof(cases "common_sudiv_exists (two_chain_vertical_boundary two_chain) \") case True show ?thesis using gen_common_subdivision_imp_eq_line_integral(2)[OF True two_chain_vertical_boundary_is_boundary_chain hv_props(6) integ_exis_vert finite_two_chain_vertical_boundary[OF assms(4)] hv_props(5) finite_j] integ_exis_horiz[of "\"] assms(3) case_prod_conv hv_props(2) hv_props(3) by fastforce next case False have i: "{j} \ Basis" using j_is_y_axis real_pair_basis by auto have ii: " \(k2, \2)\two_chain_vertical_boundary two_chain. \b\{j}. continuous_on (path_image \2) (\x. F x \ b)" using F_anal_valid field_cont_on_typeII_region_cont_on_edges valid_typeII_div by (fastforce simp add: analytically_valid_def two_chain_vertical_boundary_def boundary_def path_image_def) show "line_integral_exists F {j} \" using common_reparam_exists_imp_eq_line_integral(2)[OF finite_j hv_props(5) finite_two_chain_vertical_boundary[OF assms(4)] hv_props(6) two_chain_vertical_boundary_is_boundary_chain ii _ hv_props(7) type_II_chain_vert_bound_valid] integ_exis_horiz[of "\"] assms(3) hv_props False by fastforce qed qed lemma GreenThm_typeII_twoChain: shows "two_chain_integral two_chain (partial_vector_derivative (\a. (F a) \ j) i) = one_chain_line_integral F {j} (two_chain_boundary two_chain)" proof (simp add: two_chain_boundary_def one_chain_line_integral_def two_chain_integral_def) let ?F_a' = "partial_vector_derivative (\a. (F a) \ j) i" have "(\(k,g)\ boundary twoCube. k * line_integral F {j} g) = integral (cubeImage twoCube) (\a. ?F_a' a)" if "twoCube \ two_chain" for twoCube using green_typeII_cube.GreenThm_typeII_twoCube(1) valid_typeII_div F_anal_valid one_chain_line_integral_def valid_two_chain_def by (simp add: R2_axioms green_typeII_cube_axioms_def green_typeII_cube_def that) then have double_sum_eq_sum: "(\twoCube\(two_chain).(\(k,g)\ boundary twoCube. k * line_integral F {j} g)) = (\twoCube\(two_chain). integral (cubeImage twoCube) (\a. ?F_a' a))" using Finite_Cartesian_Product.sum_cong_aux by auto have pairwise_disjoint_boundaries: "\x\ (boundary ` two_chain). (\y\ (boundary ` two_chain). (x \ y \ (x \ y = {})))" using valid_typeII_div by (fastforce simp add: image_def valid_two_chain_def pairwise_def) have finite_boundaries: "\B \ (boundary` two_chain). finite B" using valid_typeII_div image_iff by (fastforce simp add: valid_two_cube_def valid_two_chain_def) have boundary_inj: "inj_on boundary two_chain" using valid_typeII_div by (force simp add: valid_two_cube_def valid_two_chain_def pairwise_def inj_on_def) have "(\x\(\x\two_chain. boundary x). case x of (k, g) \ k * line_integral F {j} g) = (\twoCube\(two_chain).(\(k,g)\ boundary twoCube. k * line_integral F {j} g))" using sum.reindex[OF boundary_inj, of "\x. (\(k,g)\ x. k * line_integral F {j} g)"] using sum.Union_disjoint[OF finite_boundaries pairwise_disjoint_boundaries, of "\x. case x of (k, g) \ (k::int) * line_integral F {j} g"] by auto then show "(\C\two_chain. integral (cubeImage C) (\a. ?F_a' a)) = (\x\(\x\two_chain. boundary x). case x of (k, g) \ k * line_integral F {j} g)" using double_sum_eq_sum by auto qed lemma GreenThm_typeII_divisible: assumes gen_division: "gen_division s (cubeImage ` two_chain)" (*This should follow from the assumption that images are not negligible*) shows "integral s (partial_vector_derivative (\x. (F x) \ j) i) = one_chain_line_integral F {j} (two_chain_boundary two_chain)" proof - let ?F_a' = "(partial_vector_derivative (\x. (F x) \ j) i)" have "integral s (\x. ?F_a' x) = two_chain_integral two_chain (\a. ?F_a' a)" proof (simp add: two_chain_integral_def) have partial_deriv_integrable: "(?F_a' has_integral (integral (cubeImage twoCube) ?F_a')) (cubeImage twoCube)" if "twoCube \ two_chain" for twoCube by (simp add: analytically_valid_imp_part_deriv_integrable_on F_anal_valid integrable_integral that) then have partial_deriv_integrable: "\twoCubeImg. twoCubeImg \ cubeImage ` two_chain \ ((\x. ?F_a' x) has_integral (integral (twoCubeImg) (\x. ?F_a' x))) (twoCubeImg)" using integrable_neg by force have finite_images: "finite (cubeImage ` two_chain)" using gen_division gen_division_def by auto have negligible_images: "pairwise (\S S'. negligible (S \ S')) (cubeImage ` two_chain)" using gen_division by (auto simp add: gen_division_def pairwise_def) have "inj_on cubeImage two_chain" using valid_typeII_div valid_two_chain_def by auto then have "(\twoCubeImg\cubeImage ` two_chain. integral twoCubeImg (\x. ?F_a' x)) = (\C\two_chain. integral (cubeImage C) (\a. ?F_a' a))" using sum.reindex by auto then show "integral s (\x. ?F_a' x) = (\C\two_chain. integral (cubeImage C) (\a. ?F_a' a))" using has_integral_Union[OF finite_images partial_deriv_integrable negligible_images] gen_division by (auto simp add: gen_division_def) qed then show ?thesis using GreenThm_typeII_twoChain F_anal_valid by auto qed lemma GreenThm_typeII_divisible_region_boundary_gen: assumes only_horizontal_division: "only_horizontal_division \ two_chain" shows "integral s (partial_vector_derivative (\x. (F x) \ j) i) = one_chain_line_integral F {j} \" proof - let ?F_a' = "(partial_vector_derivative (\x. (F x) \ j) i)" (*Proving that line_integral on the x axis is 0 for the vertical 1-cubes*) have horiz_line_integral_zero: "one_chain_line_integral F {j} (two_chain_horizontal_boundary two_chain) = 0" proof (simp add: one_chain_line_integral_def) have "line_integral F {j} (snd oneCube) = 0" if "oneCube \ two_chain_horizontal_boundary(two_chain)" for oneCube proof - from that obtain x y1 y2 k where horiz_edge_def: "oneCube = (k, (\t::real. ((1 - t) * (y2) + t * y1, x::real)))" using valid_typeII_div by (auto simp add: typeII_twoCube_def two_chain_horizontal_boundary_def horizontal_boundary_def) let ?horiz_edge = "(snd oneCube)" have horiz_edge_y_const: "\t. (?horiz_edge t) \ j = x" by (auto simp add: horiz_edge_def real_inner_class.inner_add_left euclidean_space_class.inner_not_same_Basis j_is_y_axis) have horiz_edge_is_straight_path: "?horiz_edge = (\t. (y2 + t * (y1 - y2), x))" by (auto simp: horiz_edge_def algebra_simps) have "\x. ?horiz_edge differentiable at x" using horiz_edge_is_straight_path straight_path_diffrentiable_y by (metis mult_commute_abs) then show "line_integral F {j} (snd oneCube) = 0" using line_integral_on_pair_straight_path(1) j_is_y_axis real_pair_basis horiz_edge_y_const by blast qed then show "(\x\two_chain_horizontal_boundary two_chain. case x of (k, g) \ k * line_integral F {j} g) = 0" by (force intro: sum.neutral) qed (*then, the x axis line_integral on the boundaries of the twoCube is equal to the line_integral on the horizontal boundaries of the twoCube \ 1*) have boundary_is_finite: "finite (two_chain_boundary two_chain)" unfolding two_chain_boundary_def proof (rule finite_UN_I) show "finite two_chain" using valid_typeII_div finite_image_iff gen_division_def valid_two_chain_def by auto show "\a. a \ two_chain \ finite (boundary a)" by (simp add: boundary_def horizontal_boundary_def vertical_boundary_def) qed have boundary_is_vert_hor: "two_chain_boundary two_chain = (two_chain_vertical_boundary two_chain) \ (two_chain_horizontal_boundary two_chain)" by(auto simp add: two_chain_boundary_def two_chain_vertical_boundary_def two_chain_horizontal_boundary_def boundary_def) then have hor_vert_finite: "finite (two_chain_vertical_boundary two_chain)" "finite (two_chain_horizontal_boundary two_chain)" using boundary_is_finite by auto have horiz_verti_disjoint: "(two_chain_vertical_boundary two_chain) \ (two_chain_horizontal_boundary two_chain) = {}" proof (simp add: two_chain_vertical_boundary_def two_chain_horizontal_boundary_def horizontal_boundary_def vertical_boundary_def) show "(\x\two_chain. {(- 1, \y. x (0, y)), (1::int, \y. x (1::real, y))}) \ (\x\two_chain. {(1, \z. x (z, 0)), (- 1, \z. x (z, 1))}) = {}" proof - have "{(- 1, \y. twoCube (0, y)), (1::int, \y. twoCube (1, y))} \ {(1, \z. twoCube2 (z, 0)), (- 1, \z. twoCube2 (z, 1))} = {}" if "twoCube\two_chain" "twoCube2\two_chain" for twoCube twoCube2 proof (cases "twoCube = twoCube2") case True have card_4: "card {(- 1, \y. twoCube2 (0::real, y)), (1::int, \y. twoCube2 (1, y)), (1, \x. twoCube2 (x, 0)), (- 1, \x. twoCube2 (x, 1))} = 4" using valid_typeII_div valid_two_chain_def that(2) by (auto simp add: boundary_def vertical_boundary_def horizontal_boundary_def valid_two_cube_def) show ?thesis using card_4 by (auto simp: True card_insert_if split: if_split_asm) next case False show ?thesis using valid_typeII_div valid_two_chain_def by (simp add: boundary_def vertical_boundary_def horizontal_boundary_def pairwise_def \twoCube \ twoCube2\ that) qed then have "\ ((\twoCube. {(- 1, \y. twoCube (0::real, y)), (1::real, \y. twoCube (1::real, y))}) ` two_chain) \ \ ((\twoCube. {(1::int, \z. twoCube (z, 0::real)), (- 1, \z. twoCube (z, 1::real))}) ` two_chain) = {}" by (fastforce simp add: Union_disjoint) then show ?thesis by force qed qed have "one_chain_line_integral F {j} (two_chain_boundary two_chain) = one_chain_line_integral F {j} (two_chain_vertical_boundary two_chain) + one_chain_line_integral F {j} (two_chain_horizontal_boundary two_chain)" using boundary_is_vert_hor horiz_verti_disjoint by (simp add: hor_vert_finite sum.union_disjoint one_chain_line_integral_def) then have y_axis_line_integral_is_only_vertical: "one_chain_line_integral F {j} (two_chain_boundary two_chain) = one_chain_line_integral F {j} (two_chain_vertical_boundary two_chain)" using horiz_line_integral_zero by auto (*Since \ \ the boundary of the 2-chain and the horizontal boundaries are \ \, then there is some \ \ \\<^sub>H(2-\) such that \ = \ \ \\<^sub>v(2-\)*) obtain \ \ where hv_props: "finite \" "(\(k,\) \ \. (\(k', \') \ two_chain_horizontal_boundary two_chain. (\a \ {0 .. 1}. \b \ {0..1}. a \ b \ subpath a b \' = \)))" "\ = \ \ \" "(common_sudiv_exists (two_chain_vertical_boundary two_chain) \ \ common_reparam_exists \ (two_chain_vertical_boundary two_chain))" "finite \" "boundary_chain \" "\(k,\)\\. valid_path \" using only_horizontal_division by(fastforce simp add: only_horizontal_division_def) have "finite {j}" by auto then have eq_integrals: " one_chain_line_integral F {j} \ = one_chain_line_integral F {j} (two_chain_vertical_boundary two_chain)" proof(cases "common_sudiv_exists (two_chain_vertical_boundary two_chain) \") case True then show ?thesis using gen_common_subdivision_imp_eq_line_integral(1)[OF True two_chain_vertical_boundary_is_boundary_chain hv_props(6) _ hor_vert_finite(1) hv_props(5)] typeII_chain_line_integral_exists_boundary' by force next case False have integ_exis_vert: "\(k,\) \ two_chain_vertical_boundary two_chain. line_integral_exists F {j} \" using typeII_chain_line_integral_exists_boundary' assms by (fastforce simp add: valid_two_chain_def) have integ_exis: "\(k,\) \ two_chain_boundary two_chain. line_integral_exists F {j} \" using typeII_cube_line_integral_exists_boundary by blast have valid_paths: "\(k,\) \ two_chain_horizontal_boundary two_chain. valid_path \" using typeII_edges_are_valid_paths valid_typeII_div by (fastforce simp add: two_chain_boundary_def two_chain_horizontal_boundary_def boundary_def) have integ_exis_horiz: "(\k \. (\(k', \')\two_chain_horizontal_boundary two_chain. \a\{0..1}. \b\{0..1}. a \ b \ subpath a b \' = \) \ line_integral_exists F {j} \)" using integ_exis valid_paths line_integral_exists_subpath[of "F" "{j}"] by (fastforce simp add: two_chain_boundary_def two_chain_horizontal_boundary_def two_chain_vertical_boundary_def boundary_def) have finite_j: "finite {j}" by auto have i: "{j} \ Basis" using j_is_y_axis real_pair_basis by auto have ii: " \(k2, \2)\two_chain_vertical_boundary two_chain. \b\{j}. continuous_on (path_image \2) (\x. F x \ b)" using valid_typeII_div field_cont_on_typeII_region_cont_on_edges F_anal_valid by (fastforce simp add: analytically_valid_def two_chain_vertical_boundary_def boundary_def path_image_def) show "one_chain_line_integral F {j} \ = one_chain_line_integral F {j} (two_chain_vertical_boundary two_chain)" using hv_props(4) False common_reparam_exists_imp_eq_line_integral(1)[OF finite_j hv_props(5) hor_vert_finite(1) hv_props(6) two_chain_vertical_boundary_is_boundary_chain ii _ hv_props(7) type_II_chain_vert_bound_valid] by fastforce qed (*Since \ \ \\<^sub>H(2-\), then the line_integral on the x axis along \ is 0, and from 1 Q.E.D. *) have line_integral_on_path: "one_chain_line_integral F {j} \ = one_chain_line_integral F {j} (two_chain_vertical_boundary two_chain)" proof (simp only: one_chain_line_integral_def) have "line_integral F {j} (snd oneCube) = 0" if oneCube: "oneCube \ \" for oneCube proof - obtain k \ where k_gamma: "(k,\) = oneCube" using prod.collapse by blast then obtain k' \' a b where kp_gammap: "(k'::int, \') \ two_chain_horizontal_boundary two_chain" "a \ {0 .. 1}" "b \ {0..1}" "subpath a b \' = \" using hv_props oneCube by (smt case_prodE split_conv) obtain x y1 y2 where horiz_edge_def: "(k',\') = (k', (\t::real. ((1 - t) * (y2) + t * y1, x::real)))" using valid_typeII_div kp_gammap by (auto simp add: typeII_twoCube_def two_chain_horizontal_boundary_def horizontal_boundary_def) have horiz_edge_y_const: "\t. \ (t) \ j = x" using horiz_edge_def kp_gammap(4) by (auto simp add: real_inner_class.inner_add_left euclidean_space_class.inner_not_same_Basis j_is_y_axis subpath_def) have horiz_edge_is_straight_path: "\ = (\t::real. ((1*y2 - a*y2) + a*y1 + ((b-a)*y1 - (b - a)*y2)*t, x::real))" proof - fix t::real have "(1 - (b - a)*t + a) * (y2) + ((b-a)*t + a) * y1 = (1 - (b - a)*t + a) * (y2) + ((b-a)*t + a) * y1" by auto then have "\ = (\t::real. ((1 - (b - a)*t - a) * (y2) + ((b-a)*t + a) * y1, x::real))" using horiz_edge_def Product_Type.snd_conv Product_Type.fst_conv kp_gammap(4) by (simp add: subpath_def diff_diff_eq[symmetric]) also have "\ = (\t::real. ((1*y2 - (b - a)*y2*t - a*y2) + ((b-a)*y1*t + a*y1), x::real))" by(auto simp add: ring_class.ring_distribs(2) Groups.diff_diff_eq left_diff_distrib) also have "... = (\t::real. ((1*y2 - a*y2) + a*y1 + ((b-a)*y1 - (b - a)*y2)*t, x::real))" by (force simp add: left_diff_distrib) finally show "\ = (\t::real. ((1*y2 - a*y2) + a*y1 + ((b-a)*y1 - (b - a)*y2)*t, x::real))" . qed show "line_integral F {j} (snd oneCube) = 0" proof - have "\x. \ differentiable at x" by (simp add: horiz_edge_is_straight_path straight_path_diffrentiable_y) then have "line_integral F {j} \ = 0" by (simp add: horiz_edge_y_const line_integral_on_pair_straight_path(1)) then show ?thesis using Product_Type.snd_conv k_gamma by auto qed qed then have "\x\\. (case x of (k, g) \ (k::int) * line_integral F {j} g) = 0" by auto then show "(\x\\. case x of (k, g) \ real_of_int k * line_integral F {j} g) = (\x\two_chain_vertical_boundary two_chain. case x of (k, g) \ real_of_int k * line_integral F {j} g)" using hv_props(1) hv_props(3) hv_props(5) sum_zero_set hor_vert_finite(1) eq_integrals by (clarsimp simp add: one_chain_line_integral_def sum_zero_set) qed then have "one_chain_line_integral F {j} \ = one_chain_line_integral F {j} (two_chain_vertical_boundary two_chain)" using line_integral_on_path by auto then have "one_chain_line_integral F {j} \ = one_chain_line_integral F {j} (two_chain_boundary two_chain)" using y_axis_line_integral_is_only_vertical by auto then show ?thesis using valid_typeII_div GreenThm_typeII_divisible by auto qed lemma GreenThm_typeII_divisible_region_boundary: assumes two_cubes_trace_vertical_boundaries: "two_chain_vertical_boundary two_chain \ \" and boundary_of_region_is_subset_of_partition_boundary: "\ \ two_chain_boundary two_chain" shows "integral s (partial_vector_derivative (\x. (F x) \ j) i) = one_chain_line_integral F {j} \" proof - let ?F_a' = "(partial_vector_derivative (\x. (F x) \ j) i)" (*Proving that line_integral on the x axis is 0 for the vertical 1-cubes*) have horiz_line_integral_zero: "one_chain_line_integral F {j} (two_chain_horizontal_boundary two_chain) = 0" proof (simp add: one_chain_line_integral_def) have "line_integral F {j} (snd oneCube) = 0" if one: "oneCube \ two_chain_horizontal_boundary(two_chain)" for oneCube proof - obtain x y1 y2 k where horiz_edge_def: "oneCube = (k, (\t::real. ((1 - t) * (y2) + t * y1, x::real)))" using valid_typeII_div one by (auto simp add: typeII_twoCube_def two_chain_horizontal_boundary_def horizontal_boundary_def) let ?horiz_edge = "(snd oneCube)" have horiz_edge_y_const: "\t. (?horiz_edge t) \ j = x" using horiz_edge_def by (auto simp add: real_inner_class.inner_add_left euclidean_space_class.inner_not_same_Basis j_is_y_axis) have horiz_edge_is_straight_path: "?horiz_edge = (\t. (y2 + t * (y1 - y2), x))" by (simp add: add_diff_eq diff_add_eq mult.commute right_diff_distrib horiz_edge_def) show "line_integral F {j} (snd oneCube) = 0" by (metis horiz_edge_is_straight_path horiz_edge_y_const line_integral_on_pair_straight_path(1) mult.commute straight_path_diffrentiable_y) qed then show "(\x\two_chain_horizontal_boundary two_chain. case x of (k, g) \ k * line_integral F {j} g) = 0" by (simp add: prod.case_eq_if) qed (*then, the x axis line_integral on the boundaries of the twoCube is equal to the line_integral on the horizontal boundaries of the twoCube \ 1*) have boundary_is_finite: "finite (two_chain_boundary two_chain)" unfolding two_chain_boundary_def proof (rule finite_UN_I) show "finite two_chain" using valid_typeII_div finite_image_iff gen_division_def valid_two_chain_def by auto show "\a. a \ two_chain \ finite (boundary a)" by (simp add: boundary_def horizontal_boundary_def vertical_boundary_def) qed have boundary_is_vert_hor: "two_chain_boundary two_chain = (two_chain_vertical_boundary two_chain) \ (two_chain_horizontal_boundary two_chain)" by(auto simp add: two_chain_boundary_def two_chain_vertical_boundary_def two_chain_horizontal_boundary_def boundary_def) then have hor_vert_finite: "finite (two_chain_vertical_boundary two_chain)" "finite (two_chain_horizontal_boundary two_chain)" using boundary_is_finite by auto have horiz_verti_disjoint: "(two_chain_vertical_boundary two_chain) \ (two_chain_horizontal_boundary two_chain) = {}" proof (simp add: two_chain_vertical_boundary_def two_chain_horizontal_boundary_def horizontal_boundary_def vertical_boundary_def) show "(\x\two_chain. {(- 1, \y. x (0, y)), (1::int, \y. x (1::real, y))}) \ (\x\two_chain. {(1, \z. x (z, 0)), (- 1, \z. x (z, 1))}) = {}" proof - have "{(- 1, \y. twoCube (0, y)), (1::int, \y. twoCube (1, y))} \ {(1, \y. twoCube2 (y, 0)), (- 1, \y. twoCube2 (y, 1))} = {}" if "twoCube\two_chain" "twoCube2\two_chain" for twoCube twoCube2 proof (cases "twoCube = twoCube2") case True have "card {(- 1, \y. twoCube2 (0::real, y)), (1::int, \y. twoCube2 (1, y)), (1, \x. twoCube2 (x, 0)), (- 1, \x. twoCube2 (x, 1))} = 4" using valid_typeII_div valid_two_chain_def that(2) by (auto simp add: boundary_def vertical_boundary_def horizontal_boundary_def valid_two_cube_def) then show ?thesis by (auto simp: True card_insert_if split: if_split_asm) next case False show ?thesis using valid_typeII_div valid_two_chain_def by (simp add: boundary_def vertical_boundary_def horizontal_boundary_def pairwise_def \twoCube \ twoCube2\ that(1) that(2)) qed then have "\ ((\twoCube. {(- 1, \y. twoCube (0::real, y)), (1::real, \y. twoCube (1::real, y))}) ` two_chain) \ \ ((\twoCube. {(1::int, \z. twoCube (z, 0::real)), (- 1, \z. twoCube (z, 1::real))}) ` two_chain) = {}" by (force simp add: Complete_Lattices.Union_disjoint) then show ?thesis by force qed qed have "one_chain_line_integral F {j} (two_chain_boundary two_chain) = one_chain_line_integral F {j} (two_chain_vertical_boundary two_chain) + one_chain_line_integral F {j} (two_chain_horizontal_boundary two_chain)" using boundary_is_vert_hor horiz_verti_disjoint by (auto simp add: one_chain_line_integral_def hor_vert_finite sum.union_disjoint) then have y_axis_line_integral_is_only_vertical: "one_chain_line_integral F {j} (two_chain_boundary two_chain) = one_chain_line_integral F {j} (two_chain_vertical_boundary two_chain)" using horiz_line_integral_zero by auto (*Since \ \ the boundary of the 2-chain and the horizontal boundaries are \ \, then there is some \ \ \\<^sub>H(2-\) such that \ = \ \ \\<^sub>v(2-\)*) have "\\. \ \ (two_chain_horizontal_boundary two_chain) \ \ = \ \ (two_chain_vertical_boundary two_chain)" proof let ?\ = "\ - (two_chain_vertical_boundary two_chain)" show "?\ \ two_chain_horizontal_boundary two_chain \ \ = ?\ \ two_chain_vertical_boundary two_chain" using two_cubes_trace_vertical_boundaries boundary_of_region_is_subset_of_partition_boundary boundary_is_vert_hor by blast qed then obtain \ where h_props: "\ \ (two_chain_horizontal_boundary two_chain)" "\ = \ \ (two_chain_vertical_boundary two_chain)" by auto have h_vert_disj: "\ \ (two_chain_vertical_boundary two_chain) = {}" using horiz_verti_disjoint h_props(1) by auto have h_finite: "finite \" using hor_vert_finite h_props(1) Finite_Set.rev_finite_subset by force have line_integral_on_path: "one_chain_line_integral F {j} \ = one_chain_line_integral F {j} \ + one_chain_line_integral F {j} (two_chain_vertical_boundary two_chain)" by (auto simp add: one_chain_line_integral_def h_props sum.union_disjoint[OF h_finite hor_vert_finite(1) h_vert_disj]) (*Since \ \ \\<^sub>H(2-\), then the line_integral on the x axis along \ is 0, and from 1 Q.E.D. *) have "one_chain_line_integral F {j} \ = 0" proof (simp add: one_chain_line_integral_def) have "line_integral F {j} (snd oneCube) = 0" if oneCube: "oneCube \ two_chain_horizontal_boundary(two_chain)" for oneCube proof - obtain x y1 y2 k where horiz_edge_def: "oneCube = (k, (\t::real. ((1 - t) * (y2) + t * y1, x::real)))" using valid_typeII_div oneCube by (auto simp add: typeII_twoCube_def two_chain_horizontal_boundary_def horizontal_boundary_def) let ?horiz_edge = "(snd oneCube)" have horiz_edge_y_const: "\t. (?horiz_edge t) \ j = x" by (simp add: j_is_y_axis horiz_edge_def) have horiz_edge_is_straight_path: "?horiz_edge = (\t. (y2 + t * (y1 - y2), x))" by (simp add: add_diff_eq diff_add_eq mult.commute right_diff_distrib horiz_edge_def) show "line_integral F {j} (snd oneCube) = 0" by (simp add: horiz_edge_is_straight_path j_is_y_axis line_integral_on_pair_straight_path(1) mult.commute straight_path_diffrentiable_y) qed then have "\oneCube \ \. line_integral F {j} (snd oneCube) = 0" using h_props by auto then have "\x\\. (case x of (k, g) \ (k::int) * line_integral F {j} g) = 0" by auto then show "(\x\\. case x of (k, g) \ k * line_integral F {j} g) = 0" by simp qed then have "one_chain_line_integral F {j} \ = one_chain_line_integral F {j} (two_chain_vertical_boundary two_chain)" using line_integral_on_path by auto then have "one_chain_line_integral F {j} \ = one_chain_line_integral F {j} (two_chain_boundary two_chain)" using y_axis_line_integral_is_only_vertical by auto then show ?thesis using valid_typeII_div GreenThm_typeII_divisible by auto qed end locale green_typeI_cube = R2 + fixes twoC F assumes two_cube: "typeI_twoCube twoC" and valid_two_cube: "valid_two_cube twoC" and f_analytically_valid: "analytically_valid (cubeImage twoC) (\x. (F x) \ i) j" begin lemma GreenThm_typeI_twoCube: shows "integral (cubeImage twoC) (\a. - partial_vector_derivative (\p. F p \ i) j a) = one_chain_line_integral F {i} (boundary twoC)" "\(k,\) \ boundary twoC. line_integral_exists F {i} \" proof - let ?bottom_edge = "(\x. twoC(x, 0))" let ?right_edge = "(\y. twoC(1, y))" let ?top_edge = "(\x. twoC(x, 1))" let ?left_edge = "(\y. twoC(0, y))" have line_integral_around_boundary: "one_chain_line_integral F {i} (boundary twoC) = line_integral F {i} ?bottom_edge + line_integral F {i} ?right_edge - line_integral F {i} ?top_edge - line_integral F {i} ?left_edge" proof (simp add: one_chain_line_integral_def horizontal_boundary_def vertical_boundary_def boundary_def) have finite1: "finite {(- 1::int, \y. twoC (0, y)), (1, \y. twoC (1, y)), (- 1, \x. twoC (x, 1))}" by auto have sum_step1: "(\(k, g)\{(- (1::int), \y. twoC (0, y)), (1, \y. twoC (1, y)), (1, \x. twoC (x, 0)), (- 1, \x. twoC (x, 1))}. k * line_integral F {i} g) = line_integral F {i} (\x. twoC (x, 0)) + (\(k, g)\{(- (1::int), \y. twoC (0, y)), (1, \y. twoC (1, y)), (- 1, \x. twoC (x, 1))}. k * line_integral F {i} g)" using sum.insert_remove [OF finite1] valid_two_cube by (auto simp: horizontal_boundary_def vertical_boundary_def boundary_def valid_two_cube_def card_insert_if split: if_split_asm) have three_distinct_edges: "card {(- 1::int, \y. twoC (0, y)), (1, \y. twoC (1, y)), (- 1, \x. twoC (x, 1))} = 3" using valid_two_cube apply (simp add: one_chain_line_integral_def horizontal_boundary_def vertical_boundary_def boundary_def valid_two_cube_def) by (auto simp: card_insert_if split: if_split_asm) have finite2: "finite {(- 1::int, \y. twoC (0, y)), (1, \y. twoC (1, y))}" by auto have sum_step2: "(\(k, g)\{(- (1::int), \y. twoC (0, y)), (1, \y. twoC (1, y)), (-1, \x. twoC (x, 1))}. k * line_integral F {i} g) = (- line_integral F {i} (\x. twoC (x, 1))) + (\(k, g)\{(- (1::int), \y. twoC (0, y)), (1, \y. twoC (1, y))}. k * line_integral F {i} g)" using sum.insert_remove [OF finite2] three_distinct_edges by (auto simp: card_insert_if split: if_split_asm) have two_distinct_edges: "card {(- 1::int, \y. twoC (0, y)), (1, \y. twoC (1, y))} = 2" using three_distinct_edges by (simp add: one_chain_line_integral_def horizontal_boundary_def vertical_boundary_def boundary_def) have finite3: "finite {(- 1::int, \y. twoC (0, y))}" by auto have sum_step3: "(\(k, g)\{(- (1::int), \y. twoC (0, y)), (1, \y. twoC (1, y))}. k * line_integral F {i} g) = (line_integral F {i} (\y. twoC (1, y))) + (\(k, g)\{(- (1::real), \y. twoC (0, y))}. k * line_integral F {i} g)" using sum.insert_remove [OF finite2] three_distinct_edges by (auto simp: card_insert_if split: if_split_asm) show "(\x\{(- 1::int, \y. twoC (0, y)), (1, \y. twoC (1, y)), (1, \x. twoC (x, 0)), (- 1, \x. twoC (x, 1))}. case x of (k, g) \ k * line_integral F {i} g) = line_integral F {i} (\x. twoC (x, 0)) + line_integral F {i} (\y. twoC (1, y)) - line_integral F {i} (\x. twoC (x, 1)) - line_integral F {i} (\y. twoC (0, y))" using sum_step1 sum_step2 sum_step3 by auto qed obtain a b g1 g2 where twoCisTypeI: "a < b" "(\x \ cbox a b. g2 x \ g1 x)" "cubeImage twoC = {(x,y). x \ cbox a b \ y \ cbox (g2 x) (g1 x)}" "twoC = (\(x, y). ((1 - x) * a + x * b, (1 - y) * g2 ((1 - x) * a + x * b) + y * g1 ((1 - x) * a + x * b)))" "g1 piecewise_C1_differentiable_on {a .. b}" "g2 piecewise_C1_differentiable_on {a .. b}" "(\x. twoC(x, 0)) = (\x. (a + (b - a) * x, g2 (a + (b - a) * x)))" "(\y. twoC(1, y)) = (\x. (b, g2 b + x *\<^sub>R (g1 b - g2 b)))" "(\x. twoC(x, 1)) = (\x. (a + (b - a) * x, g1 (a + (b - a) * x)))" "(\y. twoC(0, y)) = (\x. (a, g2 a + x *\<^sub>R (g1 a - g2 a)))" using two_cube and typeI_cube_explicit_spec[of"twoC"] by auto have bottom_edge_smooth: "(\x. twoC (x, 0)) piecewise_C1_differentiable_on {0..1}" using typeI_twoCube_smooth_edges two_cube boundary_def vertical_boundary_def horizontal_boundary_def by auto have top_edge_smooth: "?top_edge piecewise_C1_differentiable_on {0..1}" using typeI_twoCube_smooth_edges two_cube boundary_def vertical_boundary_def horizontal_boundary_def by auto show "integral (cubeImage twoC) (\a. - partial_vector_derivative (\p. F p \ i) j a) = one_chain_line_integral F {i} (boundary twoC)" using Greens_thm_type_I[OF twoCisTypeI(3) twoCisTypeI(7) bottom_edge_smooth twoCisTypeI(8) twoCisTypeI(9) top_edge_smooth twoCisTypeI(10) f_analytically_valid twoCisTypeI(2) twoCisTypeI(1)] line_integral_around_boundary by auto have "line_integral_exists F {i} (\y. twoC (0, y))" "line_integral_exists F {i} (\x. twoC (x, 1))" "line_integral_exists F {i} (\y. twoC (1, y))" "line_integral_exists F {i} (\x. twoC (x, 0))" using Greens_thm_type_I[OF twoCisTypeI(3) twoCisTypeI(7) bottom_edge_smooth twoCisTypeI(8) twoCisTypeI(9) top_edge_smooth twoCisTypeI(10) f_analytically_valid twoCisTypeI(2) twoCisTypeI(1)] line_integral_around_boundary by auto then have "line_integral_exists F {i} \" if "(k,\) \ boundary twoC" for k \ using that by (auto simp add: boundary_def horizontal_boundary_def vertical_boundary_def) then show "\(k,\) \ boundary twoC. line_integral_exists F {i} \" by auto qed lemma line_integral_exists_on_typeI_Cube_boundaries': assumes "(k,\) \ boundary twoC" shows "line_integral_exists F {i} \" using assms two_cube valid_two_cube f_analytically_valid GreenThm_typeI_twoCube(2) by blast end locale green_typeI_chain = R2 + fixes F two_chain s assumes valid_typeI_div: "valid_typeI_division s two_chain" and F_anal_valid: "\twoC \ two_chain. analytically_valid (cubeImage twoC) (\x. (F x) \ i) j" begin lemma two_chain_valid_valid_cubes: "\two_cube \ two_chain. valid_two_cube two_cube" using valid_typeI_div by (auto simp add: valid_two_chain_def) lemma typeI_cube_line_integral_exists_boundary': assumes "\two_cube \ two_chain. typeI_twoCube two_cube" assumes "\twoC \ two_chain. analytically_valid (cubeImage twoC) (\x. (F x) \ i) j" assumes "\two_cube \ two_chain. valid_two_cube two_cube" shows "\(k,\) \ two_chain_vertical_boundary two_chain. line_integral_exists F {i} \" proof - have integ_exis: "\(k,\) \ two_chain_boundary two_chain. line_integral_exists F {i} \" using green_typeI_cube.line_integral_exists_on_typeI_Cube_boundaries'[of i j] assms using R2_axioms green_typeI_cube_axioms_def green_typeI_cube_def two_chain_boundary_def by fastforce then show integ_exis_vert: "\(k,\) \ two_chain_vertical_boundary two_chain. line_integral_exists F {i} \" by (simp add: two_chain_boundary_def two_chain_vertical_boundary_def boundary_def) qed lemma typeI_cube_line_integral_exists_boundary'': "\(k,\) \ two_chain_horizontal_boundary two_chain. line_integral_exists F {i} \" proof - have integ_exis: "\(k,\) \ two_chain_boundary two_chain. line_integral_exists F {i} \" using green_typeI_cube.line_integral_exists_on_typeI_Cube_boundaries'[of i j] valid_typeI_div apply (simp add: two_chain_boundary_def boundary_def) using F_anal_valid R2_axioms green_typeI_cube_axioms_def green_typeI_cube_def two_chain_valid_valid_cubes by fastforce then show integ_exis_horiz: "\(k,\) \ two_chain_horizontal_boundary two_chain. line_integral_exists F {i} \" by (simp add: two_chain_boundary_def two_chain_horizontal_boundary_def boundary_def) qed lemma typeI_cube_line_integral_exists_boundary: "\(k,\) \ two_chain_boundary two_chain. line_integral_exists F {i} \" using typeI_cube_line_integral_exists_boundary' typeI_cube_line_integral_exists_boundary'' apply (auto simp add: two_chain_boundary_def two_chain_horizontal_boundary_def two_chain_vertical_boundary_def) by (meson R2_axioms green_typeI_chain.F_anal_valid green_typeI_chain_axioms green_typeI_cube.line_integral_exists_on_typeI_Cube_boundaries' green_typeI_cube_axioms_def green_typeI_cube_def two_chain_valid_valid_cubes valid_typeI_div) lemma type_I_chain_horiz_bound_valid: "\(k,\) \ two_chain_horizontal_boundary two_chain. valid_path \" using typeI_edges_are_valid_paths valid_typeI_div by (force simp add: two_chain_boundary_def two_chain_horizontal_boundary_def boundary_def) lemma type_I_chain_vert_bound_valid: (*This and the previous one need to be used in all proofs*) assumes "\two_cube \ two_chain. typeI_twoCube two_cube" shows "\(k,\) \ two_chain_vertical_boundary two_chain. valid_path \" using typeI_edges_are_valid_paths assms(1) by (force simp add: two_chain_boundary_def two_chain_vertical_boundary_def boundary_def) lemma members_of_only_vertical_div_line_integrable': assumes "only_vertical_division one_chain two_chain" "(k::int, \)\one_chain" "(k::int, \)\one_chain" "finite two_chain" shows "line_integral_exists F {i} \" proof - have integ_exis: "\(k,\) \ two_chain_boundary two_chain. line_integral_exists F {i} \" using typeI_cube_line_integral_exists_boundary by blast have integ_exis_horiz: "\(k,\) \ two_chain_horizontal_boundary two_chain. line_integral_exists F {i} \" using typeI_cube_line_integral_exists_boundary'' assms by auto have valid_paths: "\(k,\) \ two_chain_vertical_boundary two_chain. valid_path \" using type_I_chain_vert_bound_valid valid_typeI_div by linarith have integ_exis_vert: "(\k \. (\(k', \')\two_chain_vertical_boundary two_chain. \a\{0..1}. \b\{0..1}. a \ b \ subpath a b \' = \) \ line_integral_exists F {i} \)" using integ_exis valid_paths line_integral_exists_subpath[of "F" "{i}"] by (fastforce simp add: two_chain_boundary_def two_chain_horizontal_boundary_def two_chain_vertical_boundary_def boundary_def) obtain \ \ where hv_props: "finite \" "(\(k,\) \ \. (\(k', \') \ two_chain_vertical_boundary two_chain. (\a \ {0 .. 1}. \b \ {0..1}. a \ b \ subpath a b \' = \)))" "one_chain = \ \ \" "(common_sudiv_exists (two_chain_horizontal_boundary two_chain) \) \ common_reparam_exists \ (two_chain_horizontal_boundary two_chain)" "finite \" "boundary_chain \" "\(k,\)\\. valid_path \" using assms(1) unfolding only_vertical_division_def by blast have finite_i: "finite {i}" by auto show "line_integral_exists F {i} \" proof(cases "common_sudiv_exists (two_chain_horizontal_boundary two_chain) \") case True show ?thesis using gen_common_subdivision_imp_eq_line_integral(2)[OF True two_chain_horizontal_boundary_is_boundary_chain hv_props(6) integ_exis_horiz finite_two_chain_horizontal_boundary[OF assms(4)] hv_props(5) finite_i] integ_exis_vert[of "\"] assms(3) case_prod_conv hv_props(2) hv_props(3) by fastforce next case False have i: "{i} \ Basis" using i_is_x_axis real_pair_basis by auto have ii: " \(k2, \2)\two_chain_horizontal_boundary two_chain. \b\{i}. continuous_on (path_image \2) (\x. F x \ b)" using assms field_cont_on_typeI_region_cont_on_edges F_anal_valid valid_typeI_div by (fastforce simp add: analytically_valid_def two_chain_horizontal_boundary_def boundary_def path_image_def) show "line_integral_exists F {i} \" using common_reparam_exists_imp_eq_line_integral(2)[OF finite_i hv_props(5) finite_two_chain_horizontal_boundary[OF assms(4)] hv_props(6) two_chain_horizontal_boundary_is_boundary_chain ii _ hv_props(7) type_I_chain_horiz_bound_valid] integ_exis_vert[of "\"] False assms(3) hv_props(2-4) by fastforce qed qed lemma GreenThm_typeI_two_chain: "two_chain_integral two_chain (\a. - partial_vector_derivative (\x. (F x) \ i) j a) = one_chain_line_integral F {i} (two_chain_boundary two_chain)" proof (simp add: two_chain_boundary_def one_chain_line_integral_def two_chain_integral_def) let ?F_b' = "partial_vector_derivative (\x. (F x) \ i) j" have all_two_cubes_have_four_distict_edges: "\twoCube \ two_chain. card (boundary twoCube) = 4" using valid_typeI_div valid_two_chain_def valid_two_cube_def by auto have no_shared_edges_have_similar_orientations: "\twoCube1 \ two_chain. \twoCube2 \ two_chain. twoCube1 \ twoCube2 \ boundary twoCube1 \ boundary twoCube2 = {}" using valid_typeI_div valid_two_chain_def by (auto simp add: pairwise_def) have "(\(k,g)\ boundary twoCube. k * line_integral F {i} g) = integral (cubeImage twoCube) (\a. - ?F_b' a)" if "twoCube \ two_chain" for twoCube proof - have analytically_val: " analytically_valid (cubeImage twoCube) (\x. F x \ i) j" using that F_anal_valid by auto show "(\(k, g)\boundary twoCube. k * line_integral F {i} g) = integral (cubeImage twoCube) (\a. - ?F_b' a)" using green_typeI_cube.GreenThm_typeI_twoCube apply (simp add: one_chain_line_integral_def) by (simp add: R2_axioms analytically_val green_typeI_cube_axioms_def green_typeI_cube_def that two_chain_valid_valid_cubes valid_typeI_div) qed then have double_sum_eq_sum: "(\twoCube\(two_chain).(\(k,g)\ boundary twoCube. k * line_integral F {i} g)) = (\twoCube\(two_chain). integral (cubeImage twoCube) (\a. - ?F_b' a))" using Finite_Cartesian_Product.sum_cong_aux by auto have pairwise_disjoint_boundaries: "\x\ (boundary ` two_chain). (\y\ (boundary ` two_chain). (x \ y \ (x \ y = {})))" using no_shared_edges_have_similar_orientations by (force simp add: image_def disjoint_iff_not_equal) have finite_boundaries: "\B \ (boundary` two_chain). finite B" using all_two_cubes_have_four_distict_edges using image_iff by fastforce have boundary_inj: "inj_on boundary two_chain" using all_two_cubes_have_four_distict_edges and no_shared_edges_have_similar_orientations by (force simp add: inj_on_def) have "(\x\(\(boundary` two_chain)). case x of (k, g) \ k * line_integral F {i} g) = (sum \ sum) (\x. case x of (k, g) \ (k::int) * line_integral F {i} g) (boundary ` two_chain)" using sum.Union_disjoint[OF finite_boundaries pairwise_disjoint_boundaries] by simp also have "... = (\twoCube\(two_chain).(\(k,g)\ boundary twoCube. k * line_integral F {i} g))" using sum.reindex[OF boundary_inj, of "\x. (\(k,g)\ x. k * line_integral F {i} g)"] by auto finally show "(\C\two_chain. - integral (cubeImage C) (partial_vector_derivative (\x. F x \ i) j)) = (\x\(\x\two_chain. boundary x). case x of (k, g) \ real_of_int k * line_integral F {i} g)" using double_sum_eq_sum by auto qed lemma GreenThm_typeI_divisible: assumes gen_division: "gen_division s (cubeImage ` two_chain)" shows "integral s (\x. - partial_vector_derivative (\a. F(a) \ i) j x) = one_chain_line_integral F {i} (two_chain_boundary two_chain)" proof - let ?F_b' = "partial_vector_derivative (\a. F(a) \ i) j" have "integral s (\x. - ?F_b' x) = two_chain_integral two_chain (\a. - ?F_b' a)" proof (simp add: two_chain_integral_def) have "(\f\two_chain. integral (cubeImage f) (partial_vector_derivative (\p. F p \ i) j)) = integral s (partial_vector_derivative (\p. F p \ i) j)" by (metis analytically_valid_imp_part_deriv_integrable_on F_anal_valid gen_division two_chain_integral_def two_chain_integral_eq_integral_divisable valid_typeI_div) then show "- integral s (partial_vector_derivative (\a. F a \ i) j) = (\C\two_chain. - integral (cubeImage C) (partial_vector_derivative (\a. F a \ i) j))" by (simp add: sum_negf) qed then show ?thesis using GreenThm_typeI_two_chain assms by auto qed lemma GreenThm_typeI_divisible_region_boundary: assumes gen_division: "gen_division s (cubeImage ` two_chain)" and two_cubes_trace_horizontal_boundaries: "two_chain_horizontal_boundary two_chain \ \" and boundary_of_region_is_subset_of_partition_boundary: "\ \ two_chain_boundary two_chain" shows "integral s (\x. - partial_vector_derivative (\a. F(a) \ i) j x) = one_chain_line_integral F {i} \" proof - let ?F_b' = "partial_vector_derivative (\a. F(a) \ i)" have all_two_cubes_have_four_distict_edges: "\twoCube \ two_chain. card (boundary twoCube) = 4" using valid_typeI_div valid_two_chain_def valid_two_cube_def by auto have no_shared_edges_have_similar_orientations: "\twoCube1 \ two_chain. \twoCube2 \ two_chain. twoCube1 \ twoCube2 \ boundary twoCube1 \ boundary twoCube2 = {}" using valid_typeI_div valid_two_chain_def by (auto simp add: pairwise_def) (*Proving that line_integral on the x axis is 0 for the vertical 1-cubes*) have vert_line_integral_zero: "one_chain_line_integral F {i} (two_chain_vertical_boundary two_chain) = 0" proof (simp add: one_chain_line_integral_def) have "line_integral F {i} (snd oneCube) = 0" if oneCube: "oneCube \ two_chain_vertical_boundary(two_chain)" for oneCube proof - obtain x y1 y2 k where vert_edge_def: "oneCube = (k, (\t::real. (x::real, (1 - t) * (y2) + t * y1)))" using valid_typeI_div oneCube by (auto simp add: typeI_twoCube_def two_chain_vertical_boundary_def vertical_boundary_def) let ?vert_edge = "(snd oneCube)" have vert_edge_x_const: "\t. (?vert_edge t) \ i = x" by (simp add: i_is_x_axis vert_edge_def) have vert_edge_is_straight_path: "?vert_edge = (\t. (x, y2 + t * (y1 - y2)))" using vert_edge_def Product_Type.snd_conv by (auto simp add: mult.commute right_diff_distrib') show ?thesis by (simp add: i_is_x_axis line_integral_on_pair_straight_path(1) mult.commute straight_path_diffrentiable_x vert_edge_is_straight_path) qed then show "(\x\two_chain_vertical_boundary two_chain. case x of (k, g) \ k * line_integral F {i} g) = 0" using comm_monoid_add_class.sum.neutral by (simp add: prod.case_eq_if) qed (*then, the x axis line_integral on the boundaries of the twoCube is equal to the line_integral on the horizontal boundaries of the twoCube \ 1*) have boundary_is_finite: "finite (two_chain_boundary two_chain)" unfolding two_chain_boundary_def by (metis all_two_cubes_have_four_distict_edges card.infinite finite_UN_I finite_imageD gen_division gen_division_def zero_neq_numeral valid_typeI_div valid_two_chain_def) have boundary_is_vert_hor: "(two_chain_boundary two_chain) = (two_chain_vertical_boundary two_chain) \ (two_chain_horizontal_boundary two_chain)" by(auto simp add: two_chain_boundary_def two_chain_vertical_boundary_def two_chain_horizontal_boundary_def boundary_def) then have hor_vert_finite: "finite (two_chain_vertical_boundary two_chain)" "finite (two_chain_horizontal_boundary two_chain)" using boundary_is_finite by auto have horiz_verti_disjoint: "(two_chain_vertical_boundary two_chain) \ (two_chain_horizontal_boundary two_chain) = {}" proof (simp add: two_chain_vertical_boundary_def two_chain_horizontal_boundary_def horizontal_boundary_def vertical_boundary_def) show "(\x\two_chain. {(- 1, \y. x (0, y)), (1::int, \y. x (1::real, y))}) \ (\x\two_chain. {(1, \z. x (z, 0)), (- 1, \z. x (z, 1))}) = {}" proof - have "{(- 1, \y. twoCube (0, y)), (1::int, \y. twoCube (1, y))} \ {(1, \z. twoCube2 (z, 0)), (- 1, \z. twoCube2 (z, 1))} = {}" if "twoCube\two_chain" "twoCube2\two_chain" for twoCube twoCube2 proof (cases "twoCube = twoCube2") case True have "card {(- 1::int, \y. twoCube2 (0::real, y)), (1::int, \y. twoCube2 (1, y)), (1, \x. twoCube2 (x, 0)), (- 1, \x. twoCube2 (x, 1))} = 4" using all_two_cubes_have_four_distict_edges that(2) by (auto simp add: boundary_def vertical_boundary_def horizontal_boundary_def) then show ?thesis by (auto simp: True card_insert_if split: if_split_asm) next case False then show ?thesis using no_shared_edges_have_similar_orientations by (simp add: that boundary_def vertical_boundary_def horizontal_boundary_def) qed then have "\ ((\twoCube. {(-1::int, \y. twoCube (0,y)), (1, \y. twoCube (1, y))}) ` two_chain) \ \ ((\twoCube. {(1, \y. twoCube (y, 0)), (-1, \z. twoCube (z, 1))}) ` two_chain) = {}" using Complete_Lattices.Union_disjoint by force then show ?thesis by force qed qed have "one_chain_line_integral F {i} (two_chain_boundary two_chain) = one_chain_line_integral F {i} (two_chain_vertical_boundary two_chain) + one_chain_line_integral F {i} (two_chain_horizontal_boundary two_chain)" using boundary_is_vert_hor horiz_verti_disjoint by (auto simp add: one_chain_line_integral_def hor_vert_finite sum.union_disjoint) then have x_axis_line_integral_is_only_horizontal: "one_chain_line_integral F {i} (two_chain_boundary two_chain) = one_chain_line_integral F {i} (two_chain_horizontal_boundary two_chain)" using vert_line_integral_zero by auto (*Since \ \ the boundary of the 2-chain and the horizontal boundaries are \ \, then there is some \ \ \\<^sub>V(2-\) such that \ = \ \ \\<^sub>H(2-\)*) have "\\. \ \ (two_chain_vertical_boundary two_chain) \ \ = \ \ (two_chain_horizontal_boundary two_chain)" proof let ?\ = "\ - (two_chain_horizontal_boundary two_chain)" show "?\ \ two_chain_vertical_boundary two_chain \ \ = ?\ \ two_chain_horizontal_boundary two_chain" using two_cubes_trace_horizontal_boundaries boundary_of_region_is_subset_of_partition_boundary boundary_is_vert_hor by blast qed then obtain \ where v_props: "\ \ (two_chain_vertical_boundary two_chain)" "\ = \ \ (two_chain_horizontal_boundary two_chain)" by auto have v_horiz_disj: "\ \ (two_chain_horizontal_boundary two_chain) = {}" using horiz_verti_disjoint v_props(1) by auto have v_finite: "finite \" using hor_vert_finite v_props(1) Finite_Set.rev_finite_subset by force have line_integral_on_path: "one_chain_line_integral F {i} \ = one_chain_line_integral F {i} \ + one_chain_line_integral F {i} (two_chain_horizontal_boundary two_chain)" by(auto simp add: one_chain_line_integral_def v_props sum.union_disjoint[OF v_finite hor_vert_finite(2) v_horiz_disj]) (*Since \ \ \\<^sub>V(2-\), then the line_integral on the x axis along \ is 0, and from 1 Q.E.D. *) have "one_chain_line_integral F {i} \ = 0" proof (simp add: one_chain_line_integral_def) have "line_integral F {i} (snd oneCube) = 0" if oneCube: "oneCube \ two_chain_vertical_boundary(two_chain)" for oneCube proof - obtain x y1 y2 k where vert_edge_def: "oneCube = (k, (\t::real. (x::real, (1 - t) * (y2) + t * y1)))" using valid_typeI_div oneCube by (auto simp add: typeI_twoCube_def two_chain_vertical_boundary_def vertical_boundary_def) let ?vert_edge = "(snd oneCube)" have vert_edge_x_const: "\t. (?vert_edge t) \ i = x" by (simp add: i_is_x_axis vert_edge_def) have vert_edge_is_straight_path: "?vert_edge = (\t. (x, y2 + t * (y1 - y2)))" by (auto simp: vert_edge_def algebra_simps) have "\x. ?vert_edge differentiable at x" by (metis mult.commute vert_edge_is_straight_path straight_path_diffrentiable_x) then show "line_integral F {i} (snd oneCube) = 0" using line_integral_on_pair_straight_path(1) vert_edge_x_const by blast qed then have "\oneCube \ \. line_integral F {i} (snd oneCube) = 0" using v_props by auto then show "(\x\\. case x of (k, g) \ k * line_integral F {i} g) = 0" using comm_monoid_add_class.sum.neutral by (simp add: prod.case_eq_if) qed then have "one_chain_line_integral F {i} \ = one_chain_line_integral F {i} (two_chain_boundary two_chain)" using x_axis_line_integral_is_only_horizontal by (simp add: line_integral_on_path) then show ?thesis using assms and GreenThm_typeI_divisible by auto qed lemma GreenThm_typeI_divisible_region_boundary_gen: assumes valid_typeI_div: "valid_typeI_division s two_chain" and f_analytically_valid: "\twoC \ two_chain. analytically_valid (cubeImage twoC) (\a. F(a) \ i) j" and only_vertical_division: "only_vertical_division \ two_chain" shows "integral s (\x. - partial_vector_derivative (\a. F(a) \ i) j x) = one_chain_line_integral F {i} \" proof - let ?F_b' = "partial_vector_derivative (\a. F(a) \ i)" have all_two_cubes_have_four_distict_edges: "\twoCube \ two_chain. card (boundary twoCube) = 4" using valid_typeI_div valid_two_chain_def valid_two_cube_def by auto have no_shared_edges_have_similar_orientations: "\twoCube1 \ two_chain. \twoCube2 \ two_chain. twoCube1 \ twoCube2 \ boundary twoCube1 \ boundary twoCube2 = {}" using valid_typeI_div valid_two_chain_def by (auto simp add: pairwise_def) (*Proving that line_integral on the x axis is 0 for the vertical 1-cubes*) have vert_line_integral_zero: "one_chain_line_integral F {i} (two_chain_vertical_boundary two_chain) = 0" proof (simp add: one_chain_line_integral_def) have "line_integral F {i} (snd oneCube) = 0" if oneCube: "oneCube \ two_chain_vertical_boundary(two_chain)" for oneCube proof - obtain x y1 y2 k where vert_edge_def: "oneCube = (k, (\t::real. (x::real, (1 - t) * (y2) + t * y1)))" using valid_typeI_div oneCube by (auto simp add: typeI_twoCube_def two_chain_vertical_boundary_def vertical_boundary_def) let ?vert_edge = "(snd oneCube)" have vert_edge_x_const: "\t. (?vert_edge t) \ i = x" by (simp add: i_is_x_axis vert_edge_def) have vert_edge_is_straight_path: "?vert_edge = (\t. (x, y2 + t * (y1 - y2)))" by (auto simp: vert_edge_def algebra_simps) show ?thesis by (simp add: i_is_x_axis line_integral_on_pair_straight_path(1) mult.commute straight_path_diffrentiable_x vert_edge_is_straight_path) qed then show "(\x\two_chain_vertical_boundary two_chain. case x of (k, g) \ k * line_integral F {i} g) = 0" using comm_monoid_add_class.sum.neutral by (simp add: prod.case_eq_if) qed (*then, the x axis line_integral on the boundaries of the twoCube is equal to the line_integral on the horizontal boundaries of the twoCube \ 1*) have boundary_is_finite: "finite (two_chain_boundary two_chain)" unfolding two_chain_boundary_def proof (rule finite_UN_I) show "finite two_chain" using assms(1) finite_imageD gen_division_def valid_two_chain_def by auto show "\a. a \ two_chain \ finite (boundary a)" by (simp add: boundary_def horizontal_boundary_def vertical_boundary_def) qed have boundary_is_vert_hor: "two_chain_boundary two_chain = (two_chain_vertical_boundary two_chain) \ (two_chain_horizontal_boundary two_chain)" by (auto simp add: two_chain_boundary_def two_chain_vertical_boundary_def two_chain_horizontal_boundary_def boundary_def) then have hor_vert_finite: "finite (two_chain_vertical_boundary two_chain)" "finite (two_chain_horizontal_boundary two_chain)" using boundary_is_finite by auto have horiz_verti_disjoint: "(two_chain_vertical_boundary two_chain) \ (two_chain_horizontal_boundary two_chain) = {}" proof (simp add: two_chain_vertical_boundary_def two_chain_horizontal_boundary_def horizontal_boundary_def vertical_boundary_def) show "(\x\two_chain. {(- 1, \y. x (0, y)), (1::int, \y. x (1::real, y))}) \ (\x\two_chain. {(1, \y. x (y, 0)), (- 1, \y. x (y, 1))}) = {}" proof - have "{(- 1, \y. twoCube (0, y)), (1::int, \y. twoCube (1, y))} \ {(1, \y. twoCube2 (y, 0)), (- 1, \y. twoCube2 (y, 1))} = {}" if "twoCube \ two_chain" "twoCube2 \ two_chain" for twoCube twoCube2 proof (cases "twoCube = twoCube2") case True have "card {(- 1::int, \y. twoCube2 (0, y)), (1::int, \y. twoCube2 (1, y)), (1, \x. twoCube2 (x, 0)), (- 1, \x. twoCube2 (x, 1))} = 4" using all_two_cubes_have_four_distict_edges that(2) by (auto simp add: boundary_def vertical_boundary_def horizontal_boundary_def) then show ?thesis by (auto simp: card_insert_if True split: if_split_asm) next case False then show ?thesis using no_shared_edges_have_similar_orientations by (simp add: that boundary_def vertical_boundary_def horizontal_boundary_def) qed then have "\ ((\twoCube. {(- 1, \y. twoCube (0, y)), (1, \y. twoCube (1, y))}) ` two_chain) \ \ ((\twoCube. {(1::int, \y. twoCube (y, 0)), (- 1, \y. twoCube (y, 1))}) ` two_chain) = {}" using Complete_Lattices.Union_disjoint by force then show ?thesis by force qed qed have "one_chain_line_integral F {i} (two_chain_boundary two_chain) = one_chain_line_integral F {i} (two_chain_vertical_boundary two_chain) + one_chain_line_integral F {i} (two_chain_horizontal_boundary two_chain)" using boundary_is_vert_hor horiz_verti_disjoint by (auto simp add: one_chain_line_integral_def hor_vert_finite sum.union_disjoint) then have x_axis_line_integral_is_only_horizontal: "one_chain_line_integral F {i} (two_chain_boundary two_chain) = one_chain_line_integral F {i} (two_chain_horizontal_boundary two_chain)" using vert_line_integral_zero by auto (*Since \ \ the boundary of the 2-chain and the horizontal boundaries are \ \, then there is some \ \ \\<^sub>V(2-\) such that \ = \ \ \\<^sub>H(2-\)*) obtain \ \ where hv_props: "finite \" "(\(k,\) \ \. (\(k', \') \ two_chain_vertical_boundary two_chain. (\a \ {0 .. 1}. \b \ {0..1}. a \ b \ subpath a b \' = \)))" "\ = \ \ \" "(common_sudiv_exists (two_chain_horizontal_boundary two_chain) \) \ common_reparam_exists \ (two_chain_horizontal_boundary two_chain)" "finite \" "boundary_chain \" "\(k,\)\\. valid_path \" using only_vertical_division by (auto simp add: only_vertical_division_def) have "finite {i}" by auto then have eq_integrals: " one_chain_line_integral F {i} \ = one_chain_line_integral F {i} (two_chain_horizontal_boundary two_chain)" proof(cases "common_sudiv_exists (two_chain_horizontal_boundary two_chain) \") case True then show ?thesis using gen_common_subdivision_imp_eq_line_integral(1)[OF True two_chain_horizontal_boundary_is_boundary_chain hv_props(6) _ hor_vert_finite(2) hv_props(5)] typeI_cube_line_integral_exists_boundary'' by force next case False have integ_exis_horiz: "\(k,\) \ two_chain_horizontal_boundary two_chain. line_integral_exists F {i} \" using typeI_cube_line_integral_exists_boundary'' assms by (fastforce simp add: valid_two_chain_def) have integ_exis: "\(k,\) \ two_chain_boundary two_chain. line_integral_exists F {i} \" using typeI_cube_line_integral_exists_boundary by blast have valid_paths: "\(k,\) \ two_chain_vertical_boundary two_chain. valid_path \" using typeI_edges_are_valid_paths assms by (fastforce simp add: two_chain_boundary_def two_chain_vertical_boundary_def boundary_def) have integ_exis_vert: "(\k \. (\(k', \') \ two_chain_vertical_boundary two_chain. \a\{0..1}. \b\{0..1}. a \ b \ subpath a b \' = \) \ line_integral_exists F {i} \)" using integ_exis valid_paths line_integral_exists_subpath[of "F" "{i}"] by (fastforce simp add: two_chain_boundary_def two_chain_vertical_boundary_def two_chain_horizontal_boundary_def boundary_def) have finite_i: "finite {i}" by auto have i: "{i} \ Basis" using i_is_x_axis real_pair_basis by auto have ii: " \(k2, \2)\two_chain_horizontal_boundary two_chain. \b\{i}. continuous_on (path_image \2) (\x. F x \ b)" using assms(1) field_cont_on_typeI_region_cont_on_edges assms(2) by (fastforce simp add: analytically_valid_def two_chain_horizontal_boundary_def boundary_def path_image_def) have *: "common_reparam_exists \ (two_chain_horizontal_boundary two_chain)" using hv_props(4) False by auto show "one_chain_line_integral F {i} \ = one_chain_line_integral F {i} (two_chain_horizontal_boundary two_chain)" using common_reparam_exists_imp_eq_line_integral(1)[OF finite_i hv_props(5) hor_vert_finite(2) hv_props(6) two_chain_horizontal_boundary_is_boundary_chain ii * hv_props(7) type_I_chain_horiz_bound_valid] by fastforce qed (*Since \ \ \\<^sub>H(2-\), then the line_integral on the x axis along \ is 0, and from 1 Q.E.D. *) have line_integral_on_path: "one_chain_line_integral F {i} \ = one_chain_line_integral F {i} (two_chain_horizontal_boundary two_chain)" proof (auto simp add: one_chain_line_integral_def) have "line_integral F {i} (snd oneCube) = 0" if oneCube: "oneCube \ \" for oneCube proof - obtain k \ where k_gamma: "(k,\) = oneCube" by (metis coeff_cube_to_path.cases) then obtain k' \' a b where kp_gammap: "(k'::int, \') \ two_chain_vertical_boundary two_chain" "a \ {0 .. 1}" "b \ {0..1}" "subpath a b \' = \" using hv_props oneCube by (smt case_prodE split_conv) obtain x y1 y2 where vert_edge_def: "(k',\') = (k', (\t::real. (x::real, (1 - t) * (y2) + t * y1)))" using valid_typeI_div kp_gammap by (auto simp add: typeI_twoCube_def two_chain_vertical_boundary_def vertical_boundary_def) have vert_edge_x_const: "\t. \ (t) \ i = x" by (metis (no_types, lifting) Pair_inject fstI i_is_x_axis inner_Pair_0(2) kp_gammap(4) real_inner_1_right subpath_def vert_edge_def) have "\ = (\t::real. (x::real, (1 - (b - a)*t - a) * (y2) + ((b-a)*t + a) * y1))" using vert_edge_def Product_Type.snd_conv Product_Type.fst_conv kp_gammap(4) by (simp add: subpath_def diff_diff_eq[symmetric]) also have "... = (\t::real. (x::real, (1*y2 - a*y2) + a*y1 + ((b-a)*y1 - (b - a)*y2)*t))" by (simp add: algebra_simps) finally have vert_edge_is_straight_path: "\ = (\t::real. (x::real, (1*y2 - a*y2) + a*y1 + ((b-a)*y1 - (b - a)*y2)*t))" . show "line_integral F {i} (snd oneCube) = 0" proof - have "\x. \ differentiable at x" by (simp add: straight_path_diffrentiable_x vert_edge_is_straight_path) then have "line_integral F {i} \ = 0" using line_integral_on_pair_straight_path(1) vert_edge_x_const by blast then show ?thesis using Product_Type.snd_conv k_gamma by auto qed qed then have "\x\\. (case x of (k, g) \ (k::int) * line_integral F {i} g) = 0" by auto then show "(\x\\. case x of (k, g) \ real_of_int k * line_integral F {i} g) = (\x\two_chain_horizontal_boundary two_chain. case x of (k, g) \ of_int k * line_integral F {i} g)" using hv_props(1) hv_props(3) hv_props(5) sum_zero_set hor_vert_finite(2) eq_integrals apply(auto simp add: one_chain_line_integral_def) by (smt Un_commute sum_zero_set) qed then have "one_chain_line_integral F {i} \ = one_chain_line_integral F {i} (two_chain_boundary two_chain)" using x_axis_line_integral_is_only_horizontal line_integral_on_path by auto then show ?thesis using assms GreenThm_typeI_divisible by auto qed end locale green_typeI_typeII_chain = R2: R2 i j + T1: green_typeI_chain i j F two_chain_typeI + T2: green_typeII_chain i j F two_chain_typeII for i j F two_chain_typeI two_chain_typeII begin lemma GreenThm_typeI_typeII_divisible_region_boundary: assumes gen_divisions: "gen_division s (cubeImage ` two_chain_typeI)" "gen_division s (cubeImage ` two_chain_typeII)" and typeI_two_cubes_trace_horizontal_boundaries: "two_chain_horizontal_boundary two_chain_typeI \ \" and typeII_two_cubes_trace_vertical_boundaries: "two_chain_vertical_boundary two_chain_typeII \ \" and boundary_of_region_is_subset_of_partition_boundaries: "\ \ two_chain_boundary two_chain_typeI" "\ \ two_chain_boundary two_chain_typeII" shows "integral s (\x. partial_vector_derivative (\a. F a \ j) i x - partial_vector_derivative (\a. F a \ i) j x) = one_chain_line_integral F {i, j} \" proof - let ?F_b' = "partial_vector_derivative (\a. F a \ i) j" let ?F_a' = "partial_vector_derivative (\a. F a \ j) i" have typeI_regions_integral: "integral s (\x. - partial_vector_derivative (\a. F a \ i) j x) = one_chain_line_integral F {i} \" using T1.GreenThm_typeI_divisible_region_boundary gen_divisions(1) typeI_two_cubes_trace_horizontal_boundaries boundary_of_region_is_subset_of_partition_boundaries(1) by blast have typeII_regions_integral: "integral s (partial_vector_derivative (\x. F x \ j) i) = one_chain_line_integral F {j} \" using T2.GreenThm_typeII_divisible_region_boundary gen_divisions(2) typeII_two_cubes_trace_vertical_boundaries boundary_of_region_is_subset_of_partition_boundaries(2) by auto have integral_dis: "integral s (\x. ?F_a' x - ?F_b' x) = integral s (\x. ?F_a' x) + integral s (\x. - ?F_b' x)" proof - have "\twoCube \ two_chain_typeII. (?F_a' has_integral integral (cubeImage twoCube) ?F_a') (cubeImage twoCube)" by (simp add: analytically_valid_imp_part_deriv_integrable_on T2.F_anal_valid has_integral_iff) then have "\u. u \ (cubeImage ` two_chain_typeII) \ (?F_a' has_integral integral u ?F_a') u" by auto then have "(?F_a' has_integral (\img\cubeImage ` two_chain_typeII. integral img ?F_a')) s" using gen_divisions(2) unfolding gen_division_def by (metis has_integral_Union) then have F_a'_integrable: "(?F_a' integrable_on s)" by auto have "\twoCube \ two_chain_typeI. (?F_b' has_integral integral (cubeImage twoCube) ?F_b') (cubeImage twoCube)" using analytically_valid_imp_part_deriv_integrable_on T1.F_anal_valid by blast then have "\u. u \ (cubeImage ` two_chain_typeI) \ (?F_b' has_integral integral u ?F_b') u" by auto then have "(?F_b' has_integral (\img\cubeImage ` two_chain_typeI. integral img ?F_b')) s" using gen_divisions(1) unfolding gen_division_def by (metis has_integral_Union) then show ?thesis by (simp add: F_a'_integrable Henstock_Kurzweil_Integration.integral_diff has_integral_iff) qed have line_integral_dist: "one_chain_line_integral F {i, j} \ = one_chain_line_integral F {i} \ + one_chain_line_integral F {j} \" proof (simp add: one_chain_line_integral_def) have "k * line_integral F {i, j} g = k * line_integral F {i} g + k * line_integral F {j} g" if kg: "(k,g) \ \" for k g proof - obtain twoCube_typeI where twoCube_typeI_props: "twoCube_typeI \ two_chain_typeI" "(k, g) \ boundary twoCube_typeI" "typeI_twoCube twoCube_typeI" "continuous_on (cubeImage twoCube_typeI) (\x. F(x) \ i)" using boundary_of_region_is_subset_of_partition_boundaries(1) two_chain_boundary_def T1.valid_typeI_div T1.F_anal_valid kg by (auto simp add: analytically_valid_def) obtain twoCube_typeII where twoCube_typeII_props: "twoCube_typeII \ two_chain_typeII" "(k, g) \ boundary twoCube_typeII" "typeII_twoCube twoCube_typeII" "continuous_on (cubeImage twoCube_typeII) (\x. F(x) \ j)" using boundary_of_region_is_subset_of_partition_boundaries(2) two_chain_boundary_def T2.valid_typeII_div kg T2.F_anal_valid by (auto simp add: analytically_valid_def) have "line_integral F {i, j} g = line_integral F {i} g + line_integral F {j} g" proof - have int_exists_i: "line_integral_exists F {i} g" using T1.typeI_cube_line_integral_exists_boundary assms kg by (auto simp add: valid_two_chain_def) have int_exists_j: "line_integral_exists F {j} g" using T2.typeII_cube_line_integral_exists_boundary assms kg by (auto simp add: valid_two_chain_def) have finite: "finite {i, j}" by auto show ?thesis using line_integral_sum_gen[OF finite int_exists_i int_exists_j] R2.i_is_x_axis R2.j_is_y_axis by auto qed then show "k * line_integral F {i, j} g = k * line_integral F {i} g + k * line_integral F {j} g" by (simp add: distrib_left) qed then have line_integral_distrib: "(\(k,g)\\. k * line_integral F {i, j} g) = (\(k,g)\\. k * line_integral F {i} g + k * line_integral F {j} g)" by (force intro: sum.cong split_cong) have "(\x. (case x of (k, g) \ (k::int) * line_integral F {i} g) + (case x of (k, g) \ (k::int) * line_integral F {j} g)) = (\x. (case x of (k, g) \ (k * line_integral F {i} g) + (k::int) * line_integral F {j} g))" using comm_monoid_add_class.sum.distrib by auto then show "(\(k, g)\\. k * line_integral F {i, j} g) = (\(k, g)\\. (k::int) * line_integral F {i} g) + (\(k, g)\\. (k::int) * line_integral F {j} g)" using comm_monoid_add_class.sum.distrib[of "(\(k, g). k * line_integral F {i} g)" "(\(k, g). k * line_integral F {j} g)" \] line_integral_distrib by presburger qed show ?thesis using integral_dis line_integral_dist typeI_regions_integral typeII_regions_integral by auto qed lemma GreenThm_typeI_typeII_divisible_region': assumes only_vertical_division: "only_vertical_division one_chain_typeI two_chain_typeI" "boundary_chain one_chain_typeI" and only_horizontal_division: "only_horizontal_division one_chain_typeII two_chain_typeII" "boundary_chain one_chain_typeII" and typeI_and_typII_one_chains_have_gen_common_subdiv: "common_sudiv_exists one_chain_typeI one_chain_typeII" shows "integral s (\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} one_chain_typeI" "integral s (\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} one_chain_typeII" proof - let ?F_b' = "partial_vector_derivative (\x. (F x) \ i) j" let ?F_a' = "partial_vector_derivative (\x. (F x) \ j) i" have one_chain_i_integrals: "one_chain_line_integral F {i} one_chain_typeI = one_chain_line_integral F {i} one_chain_typeII \ (\(k,\)\one_chain_typeI. line_integral_exists F {i} \) \ (\(k,\)\one_chain_typeII. line_integral_exists F {i} \)" proof (intro conjI) have "finite two_chain_typeI" using T1.valid_typeI_div finite_image_iff by (auto simp add: gen_division_def valid_two_chain_def) then show ii: "\(k, \)\one_chain_typeI. line_integral_exists F {i} \" using T1.members_of_only_vertical_div_line_integrable' assms by fastforce have "finite (two_chain_horizontal_boundary two_chain_typeI)" by (meson T1.valid_typeI_div finite_imageD finite_two_chain_horizontal_boundary gen_division_def valid_two_chain_def) then have "finite one_chain_typeI" using only_vertical_division(1) only_vertical_division_def by auto moreover have "finite one_chain_typeII" using only_horizontal_division(1) only_horizontal_division_def by auto ultimately show "one_chain_line_integral F {i} one_chain_typeI = one_chain_line_integral F {i} one_chain_typeII" and "\(k, \)\one_chain_typeII. line_integral_exists F {i} \" using gen_common_subdivision_imp_eq_line_integral[OF typeI_and_typII_one_chains_have_gen_common_subdiv only_vertical_division(2) only_horizontal_division(2)] ii by auto qed have one_chain_j_integrals: "one_chain_line_integral F {j} one_chain_typeII = one_chain_line_integral F {j} one_chain_typeI \ (\(k,\)\one_chain_typeII. line_integral_exists F {j} \) \ (\(k,\)\one_chain_typeI. line_integral_exists F {j} \)" proof (intro conjI) have "finite two_chain_typeII" using T2.valid_typeII_div finite_image_iff by (auto simp add: gen_division_def valid_two_chain_def) then show ii: "\(k,\)\one_chain_typeII. line_integral_exists F {j} \" using T2.members_of_only_horiz_div_line_integrable' assms T2.two_chain_valid_valid_cubes by blast have typeII_and_typI_one_chains_have_common_subdiv: "common_sudiv_exists one_chain_typeII one_chain_typeI" by (simp add: common_sudiv_exists_comm typeI_and_typII_one_chains_have_gen_common_subdiv) have iv: "finite one_chain_typeI" using only_vertical_division(1) only_vertical_division_def by auto moreover have iv': "finite one_chain_typeII" using only_horizontal_division(1) only_horizontal_division_def by auto ultimately show "one_chain_line_integral F {j} one_chain_typeII = one_chain_line_integral F {j} one_chain_typeI" "\(k, \)\one_chain_typeI. line_integral_exists F {j} \" using gen_common_subdivision_imp_eq_line_integral[OF typeII_and_typI_one_chains_have_common_subdiv only_horizontal_division(2) only_vertical_division(2) ii] ii by auto qed have typeI_regions_integral: "integral s (\x. - ?F_b' x) = one_chain_line_integral F {i} one_chain_typeI" using T1.GreenThm_typeI_divisible_region_boundary_gen T1.valid_typeI_div T1.F_anal_valid only_vertical_division(1) by auto have typeII_regions_integral: "integral s ?F_a' = one_chain_line_integral F {j} one_chain_typeII" using T2.GreenThm_typeII_divisible_region_boundary_gen T2.valid_typeII_div T2.F_anal_valid only_horizontal_division(1) by auto have line_integral_dist: "one_chain_line_integral F {i, j} one_chain_typeI = one_chain_line_integral F {i} one_chain_typeI + one_chain_line_integral F {j} one_chain_typeI \ one_chain_line_integral F {i, j} one_chain_typeII = one_chain_line_integral F {i} one_chain_typeII + one_chain_line_integral F {j} one_chain_typeII" proof (simp add: one_chain_line_integral_def) have line_integral_distrib: "(\(k,g)\one_chain_typeI. k * line_integral F {i, j} g) = (\(k,g)\one_chain_typeI. k * line_integral F {i} g + k * line_integral F {j} g) \ (\(k,g)\one_chain_typeII. k * line_integral F {i, j} g) = (\(k,g)\one_chain_typeII. k * line_integral F {i} g + k * line_integral F {j} g)" proof - have 0: "k * line_integral F {i, j} g = k * line_integral F {i} g + k * line_integral F {j} g" if "(k,g) \ one_chain_typeII" for k g proof - have "line_integral_exists F {i} g" "line_integral_exists F {j} g" "finite {i, j}" using one_chain_i_integrals one_chain_j_integrals that by fastforce+ moreover have "{i} \ {j} = {}" by (simp add: R2.i_is_x_axis R2.j_is_y_axis) ultimately have "line_integral F {i, j} g = line_integral F {i} g + line_integral F {j} g" by (metis insert_is_Un line_integral_sum_gen(1)) then show "k * line_integral F {i, j} g = k * line_integral F {i} g + k * line_integral F {j} g" by (simp add: distrib_left) qed have "k * line_integral F {i, j} g = k * line_integral F {i} g + k * line_integral F {j} g" if "(k,g) \ one_chain_typeI" for k g proof - have "line_integral F {i, j} g = line_integral F {i} g + line_integral F {j} g" by (smt that disjoint_insert(2) finite.emptyI finite.insertI R2.i_is_x_axis inf_bot_right insert_absorb insert_commute insert_is_Un R2.j_is_y_axis line_integral_sum_gen(1) one_chain_i_integrals one_chain_j_integrals prod.case_eq_if singleton_inject snd_conv) then show "k * line_integral F {i, j} g = k * line_integral F {i} g + k * line_integral F {j} g" by (simp add: distrib_left) qed then show ?thesis using 0 by (smt sum.cong split_cong) qed show "(\(k::int, g)\one_chain_typeI. k * line_integral F {i, j} g) = (\(k, g)\one_chain_typeI. k * line_integral F {i} g) + (\(k::int, g)\one_chain_typeI. k * line_integral F {j} g) \ (\(k::int, g)\one_chain_typeII. k * line_integral F {i, j} g) = (\(k, g)\one_chain_typeII. k * line_integral F {i} g) + (\(k::int, g)\one_chain_typeII. k * line_integral F {j} g)" proof - have 0: "(\x. (case x of (k::int, g) \ k * line_integral F {i} g) + (case x of (k::int, g) \ k * line_integral F {j} g)) = (\x. (case x of (k::int, g) \ (k * line_integral F {i} g) + k * line_integral F {j} g))" using comm_monoid_add_class.sum.distrib by auto then have 1: "(\x\one_chain_typeI. (case x of (k::int, g) \ k * line_integral F {i} g) + (case x of (k::int, g) \ k * line_integral F {j} g)) = (\x\one_chain_typeI. (case x of (k::int, g) \( k * line_integral F {i} g + k * line_integral F {j} g)))" by presburger have "(\x\one_chain_typeII. (case x of (k, g) \ k * line_integral F {i} g) + (case x of (k, g) \ k * line_integral F {j} g)) = (\x\one_chain_typeII. (case x of (k, g) \( k * line_integral F {i} g + k * line_integral F {j} g)))" using 0 by presburger then show ?thesis using sum.distrib[of "(\(k, g). k * line_integral F {i} g)" "(\(k, g). k * line_integral F {j} g)" "one_chain_typeI"] sum.distrib[of "(\(k, g). k * line_integral F {i} g)" "(\(k, g). k * line_integral F {j} g)" "one_chain_typeII"] line_integral_distrib 1 by auto qed qed have integral_dis: "integral s (\x. ?F_a' x - ?F_b' x) = integral s (\x. ?F_a' x) + integral s (\x. - ?F_b' x)" proof - have "(?F_a' has_integral integral (cubeImage twoCube) ?F_a') (cubeImage twoCube)" if "twoCube \ two_chain_typeII" for twoCube by (simp add: analytically_valid_imp_part_deriv_integrable_on T2.F_anal_valid has_integral_integrable_integral that) then have "\u. u \ (cubeImage ` two_chain_typeII) \ (?F_a' has_integral integral u ?F_a') u" by auto then have "(?F_a' has_integral (\img\cubeImage ` two_chain_typeII. integral img ?F_a')) s" using T2.valid_typeII_div unfolding gen_division_def by (metis has_integral_Union) then have F_a'_integrable: "(?F_a' integrable_on s)" by auto have "\twoCube \ two_chain_typeI. (?F_b' has_integral integral (cubeImage twoCube) ?F_b') (cubeImage twoCube)" using analytically_valid_imp_part_deriv_integrable_on T1.F_anal_valid by blast then have "\u. u \ (cubeImage ` two_chain_typeI) \ (?F_b' has_integral integral u ?F_b') u" by auto then have "(?F_b' has_integral (\img\cubeImage ` two_chain_typeI. integral img ?F_b')) s" using T1.valid_typeI_div unfolding gen_division_def by (metis has_integral_Union) then show ?thesis by (simp add: F_a'_integrable Henstock_Kurzweil_Integration.integral_diff has_integral_iff) qed show "integral s (\x. ?F_a' x - ?F_b' x) = one_chain_line_integral F {i, j} one_chain_typeI" using one_chain_j_integrals integral_dis line_integral_dist typeI_regions_integral typeII_regions_integral by auto show "integral s (\x. ?F_a' x - ?F_b' x) = one_chain_line_integral F {i, j} one_chain_typeII" using one_chain_i_integrals integral_dis line_integral_dist typeI_regions_integral typeII_regions_integral by auto qed lemma GreenThm_typeI_typeII_divisible_region: assumes only_vertical_division: "only_vertical_division one_chain_typeI two_chain_typeI" "boundary_chain one_chain_typeI" and only_horizontal_division: "only_horizontal_division one_chain_typeII two_chain_typeII" "boundary_chain one_chain_typeII" and typeI_and_typII_one_chains_have_common_subdiv: "common_boundary_sudivision_exists one_chain_typeI one_chain_typeII" shows "integral s (\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} one_chain_typeI" "integral s (\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} one_chain_typeII" using GreenThm_typeI_typeII_divisible_region' only_vertical_division only_horizontal_division common_subdiv_imp_gen_common_subdiv[OF typeI_and_typII_one_chains_have_common_subdiv] by auto lemma GreenThm_typeI_typeII_divisible_region_finite_holes: assumes valid_cube_boundary: "\(k,\)\boundary C. valid_path \" and only_vertical_division: "only_vertical_division (boundary C) two_chain_typeI" and only_horizontal_division: "only_horizontal_division (boundary C) two_chain_typeII" and s_is_oneCube: "s = cubeImage C" shows "integral (cubeImage C) (\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 C)" using GreenThm_typeI_typeII_divisible_region[OF only_vertical_division two_cube_boundary_is_boundary only_horizontal_division two_cube_boundary_is_boundary common_boundary_subdiv_exists_refl[OF assms(1)]] s_is_oneCube by auto lemma GreenThm_typeI_typeII_divisible_region_equivallent_boundary: assumes gen_divisions: "gen_division s (cubeImage ` two_chain_typeI)" "gen_division s (cubeImage ` two_chain_typeII)" and typeI_two_cubes_trace_horizontal_boundaries: "two_chain_horizontal_boundary two_chain_typeI \ one_chain_typeI" and typeII_two_cubes_trace_vertical_boundaries: "two_chain_vertical_boundary two_chain_typeII \ one_chain_typeII" and boundary_of_region_is_subset_of_partition_boundaries: "one_chain_typeI \ two_chain_boundary two_chain_typeI" "one_chain_typeII \ two_chain_boundary two_chain_typeII" and typeI_and_typII_one_chains_have_common_subdiv: "common_boundary_sudivision_exists one_chain_typeI one_chain_typeII" shows "integral s (\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} one_chain_typeI" "integral s (\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} one_chain_typeII" proof - let ?F_b' = "partial_vector_derivative (\x. (F x) \ i) j" let ?F_a' = "partial_vector_derivative (\x. (F x) \ j) i" have one_chain_i_integrals: "one_chain_line_integral F {i} one_chain_typeI = one_chain_line_integral F {i} one_chain_typeII \ (\(k,\)\one_chain_typeI. line_integral_exists F {i} \) \ (\(k,\)\one_chain_typeII. line_integral_exists F {i} \)" proof (intro conjI) have i: "boundary_chain one_chain_typeI" using two_chain_boundary_is_boundary_chain boundary_chain_def boundary_of_region_is_subset_of_partition_boundaries(1) by blast have i': "boundary_chain one_chain_typeII" using two_chain_boundary_is_boundary_chain boundary_chain_def boundary_of_region_is_subset_of_partition_boundaries(2) by blast have "\k \. (k,\)\one_chain_typeI \ line_integral_exists F {i} \" using T1.typeI_cube_line_integral_exists_boundary assms by (fastforce simp add: valid_two_chain_def) then show ii: "\(k,\)\one_chain_typeI. line_integral_exists F {i} \" by auto have "finite (two_chain_boundary two_chain_typeI)" unfolding two_chain_boundary_def proof (rule finite_UN_I) show "finite two_chain_typeI" using T1.valid_typeI_div finite_imageD gen_division_def valid_two_chain_def by auto show "\a. a \ two_chain_typeI \ finite (boundary a)" by (simp add: boundary_def horizontal_boundary_def vertical_boundary_def) qed then have "finite one_chain_typeI" using boundary_of_region_is_subset_of_partition_boundaries(1) finite_subset by fastforce moreover have "finite (two_chain_boundary two_chain_typeII)" unfolding two_chain_boundary_def proof (rule finite_UN_I) show "finite two_chain_typeII" using T2.valid_typeII_div finite_imageD gen_division_def valid_two_chain_def by auto show "\a. a \ two_chain_typeII \ finite (boundary a)" by (simp add: boundary_def horizontal_boundary_def vertical_boundary_def) qed then have "finite one_chain_typeII" using boundary_of_region_is_subset_of_partition_boundaries(2) finite_subset by fastforce ultimately show "one_chain_line_integral F {i} one_chain_typeI = one_chain_line_integral F {i} one_chain_typeII" "\(k, \)\one_chain_typeII. line_integral_exists F {i} \" using ii common_subdivision_imp_eq_line_integral[OF typeI_and_typII_one_chains_have_common_subdiv i i' ii] by auto qed have one_chain_j_integrals: "one_chain_line_integral F {j} one_chain_typeI = one_chain_line_integral F {j} one_chain_typeII \ (\(k,\)\one_chain_typeI. line_integral_exists F {j} \) \ (\(k,\)\one_chain_typeII. line_integral_exists F {j} \)" proof (intro conjI) have i: "boundary_chain one_chain_typeI" and i': "boundary_chain one_chain_typeII" using two_chain_boundary_is_boundary_chain boundary_of_region_is_subset_of_partition_boundaries unfolding boundary_chain_def by blast+ have "line_integral_exists F {j} \" if "(k,\)\one_chain_typeII" for k \ proof - have F_is_continuous: "\twoC \ two_chain_typeII. continuous_on (cubeImage twoC) (\a. F(a) \ j)" using T2.F_anal_valid by(simp add: analytically_valid_def) show "line_integral_exists F {j} \" using that T2.valid_typeII_div boundary_of_region_is_subset_of_partition_boundaries(2) using green_typeII_cube.line_integral_exists_on_typeII_Cube_boundaries' assms valid_two_chain_def apply (simp add: two_chain_boundary_def) by (metis T2.typeII_cube_line_integral_exists_boundary case_prodD subset_iff that two_chain_boundary_def) qed then show ii: " \(k,\)\one_chain_typeII. line_integral_exists F {j} \" by auto have "finite (two_chain_boundary two_chain_typeI)" unfolding two_chain_boundary_def proof (rule finite_UN_I) show "finite two_chain_typeI" using T1.valid_typeI_div finite_imageD gen_division_def valid_two_chain_def by auto show "\a. a \ two_chain_typeI \ finite (boundary a)" by (simp add: boundary_def horizontal_boundary_def vertical_boundary_def) qed then have iv: "finite one_chain_typeI" using boundary_of_region_is_subset_of_partition_boundaries(1) finite_subset by fastforce have "finite (two_chain_boundary two_chain_typeII)" unfolding two_chain_boundary_def proof (rule finite_UN_I) show "finite two_chain_typeII" using T2.valid_typeII_div finite_imageD gen_division_def valid_two_chain_def by auto show "\a. a \ two_chain_typeII \ finite (boundary a)" by (simp add: boundary_def horizontal_boundary_def vertical_boundary_def) qed then have iv': "finite one_chain_typeII" using boundary_of_region_is_subset_of_partition_boundaries(2) finite_subset by fastforce have typeII_and_typI_one_chains_have_common_subdiv: "common_boundary_sudivision_exists one_chain_typeII one_chain_typeI" using typeI_and_typII_one_chains_have_common_subdiv common_boundary_sudivision_commutative by auto show "one_chain_line_integral F {j} one_chain_typeI = one_chain_line_integral F {j} one_chain_typeII" "\(k, \)\one_chain_typeI. line_integral_exists F {j} \" using common_subdivision_imp_eq_line_integral[OF typeII_and_typI_one_chains_have_common_subdiv i' i ii iv' iv] ii by auto qed have typeI_regions_integral: "integral s (\x. - ?F_b' x) = one_chain_line_integral F {i} one_chain_typeI" using T1.GreenThm_typeI_divisible_region_boundary gen_divisions(1) typeI_two_cubes_trace_horizontal_boundaries boundary_of_region_is_subset_of_partition_boundaries(1) by auto have typeII_regions_integral: "integral s ?F_a' = one_chain_line_integral F {j} one_chain_typeII" using T2.GreenThm_typeII_divisible_region_boundary gen_divisions(2) typeII_two_cubes_trace_vertical_boundaries boundary_of_region_is_subset_of_partition_boundaries(2) by auto have line_integral_dist: "one_chain_line_integral F {i, j} one_chain_typeI = one_chain_line_integral F {i} one_chain_typeI + one_chain_line_integral F {j} one_chain_typeI \ one_chain_line_integral F {i, j} one_chain_typeII = one_chain_line_integral F {i} one_chain_typeII + one_chain_line_integral F {j} one_chain_typeII" proof (simp add: one_chain_line_integral_def) have line_integral_distrib: "(\(k,g)\one_chain_typeI. k * line_integral F {i, j} g) = (\(k,g)\one_chain_typeI. k * line_integral F {i} g + k * line_integral F {j} g) \ (\(k,g)\one_chain_typeII. k * line_integral F {i, j} g) = (\(k,g)\one_chain_typeII. k * line_integral F {i} g + k * line_integral F {j} g)" proof - have 0: "k * line_integral F {i, j} g = k * line_integral F {i} g + k * line_integral F {j} g" if "(k,g) \ one_chain_typeII" for k g proof - have "line_integral F {i, j} g = line_integral F {i} g + line_integral F {j} g" proof - have finite: "finite {i, j}" by auto have line_integral_all: "\i\{i, j}. line_integral_exists F {i} g" using one_chain_i_integrals one_chain_j_integrals that by auto show ?thesis using line_integral_sum_gen[OF finite] R2.i_is_x_axis R2.j_is_y_axis line_integral_all by auto qed then show "k * line_integral F {i, j} g = k * line_integral F {i} g + k * line_integral F {j} g" by (simp add: distrib_left) qed have "k * line_integral F {i, j} g = k * line_integral F {i} g + k * line_integral F {j} g" if "(k,g) \ one_chain_typeI" for k g proof - have finite: "finite {i, j}" by auto have line_integral_all: "\i\{i, j}. line_integral_exists F {i} g" using one_chain_i_integrals one_chain_j_integrals that by auto have "line_integral F {i, j} g = line_integral F {i} g + line_integral F {j} g" using line_integral_sum_gen[OF finite] R2.i_is_x_axis R2.j_is_y_axis line_integral_all by auto then show "k * line_integral F {i, j} g = k * line_integral F {i} g + k * line_integral F {j} g" by (simp add: distrib_left) qed then show ?thesis using 0 by (smt sum.cong split_cong) qed show "(\(k::int, g)\one_chain_typeI. k * line_integral F {i, j} g) = (\(k, g)\one_chain_typeI. k * line_integral F {i} g) + (\(k::int, g)\one_chain_typeI. k * line_integral F {j} g) \ (\(k::int, g)\one_chain_typeII. k * line_integral F {i, j} g) = (\(k, g)\one_chain_typeII. k * line_integral F {i} g) + (\(k::int, g)\one_chain_typeII. k * line_integral F {j} g)" proof - have 0: "(\x. (case x of (k::int, g) \ k * line_integral F {i} g) + (case x of (k::int, g) \ k * line_integral F {j} g)) = (\x. (case x of (k::int, g) \ (k * line_integral F {i} g) + k * line_integral F {j} g))" using comm_monoid_add_class.sum.distrib by auto then have 1: "(\x\one_chain_typeI. (case x of (k::int, g) \ k * line_integral F {i} g) + (case x of (k::int, g) \ k * line_integral F {j} g)) = (\x\one_chain_typeI. (case x of (k::int, g) \( k * line_integral F {i} g + k * line_integral F {j} g)))" by presburger have "(\x\one_chain_typeII. (case x of (k, g) \ k * line_integral F {i} g) + (case x of (k, g) \ k * line_integral F {j} g)) = (\x\one_chain_typeII. (case x of (k, g) \( k * line_integral F {i} g + k * line_integral F {j} g)))" using 0 by presburger then show ?thesis using sum.distrib[of "(\(k, g). k * line_integral F {i} g)" "(\(k, g). k * line_integral F {j} g)" "one_chain_typeI"] sum.distrib[of "(\(k, g). k * line_integral F {i} g)" "(\(k, g). k * line_integral F {j} g)" "one_chain_typeII"] line_integral_distrib 1 by auto qed qed have integral_dis: "integral s (\x. ?F_a' x - ?F_b' x) = integral s (\x. ?F_a' x) + integral s (\x. - ?F_b' x)" proof - have "(?F_a' has_integral (\img\cubeImage ` two_chain_typeII. integral img ?F_a')) s" proof - have "(?F_a' has_integral integral (cubeImage twoCube) ?F_a') (cubeImage twoCube)" if "twoCube \ two_chain_typeII" for twoCube by (simp add: analytically_valid_imp_part_deriv_integrable_on T2.F_anal_valid has_integral_integrable_integral that) then have "\u. u \ (cubeImage ` two_chain_typeII) \ (?F_a' has_integral integral u ?F_a') u" by auto then show ?thesis using gen_divisions(2) unfolding gen_division_def by (metis has_integral_Union) qed then have F_a'_integrable: "(?F_a' integrable_on s)" by auto have "(?F_b' has_integral (\img\cubeImage ` two_chain_typeI. integral img ?F_b')) s" proof - have "\twoCube \ two_chain_typeI. (?F_b' has_integral integral (cubeImage twoCube) ?F_b') (cubeImage twoCube)" by (simp add: analytically_valid_imp_part_deriv_integrable_on T1.F_anal_valid has_integral_integrable_integral) then have "\u. u \ (cubeImage ` two_chain_typeI) \ (?F_b' has_integral integral u ?F_b') u" by auto then show ?thesis using gen_divisions(1) unfolding gen_division_def by (metis has_integral_Union) qed then show ?thesis using F_a'_integrable Henstock_Kurzweil_Integration.integral_diff by auto qed show "integral s (\x. ?F_a' x - ?F_b' x) = one_chain_line_integral F {i, j} one_chain_typeI" using one_chain_j_integrals integral_dis line_integral_dist typeI_regions_integral typeII_regions_integral by auto show "integral s (\x. ?F_a' x - ?F_b' x) = one_chain_line_integral F {i, j} one_chain_typeII" using one_chain_i_integrals integral_dis line_integral_dist typeI_regions_integral typeII_regions_integral by auto qed end end diff --git a/thys/Green/Paths.thy b/thys/Green/Paths.thy --- a/thys/Green/Paths.thy +++ b/thys/Green/Paths.thy @@ -1,3584 +1,3584 @@ theory Paths imports Derivs General_Utils Integrals begin (*This has everything related to paths purely*) lemma reverse_subpaths_join: shows " subpath 1 (1 / 2) p +++ subpath (1 / 2) 0 p = reversepath p" using reversepath_subpath join_subpaths_middle pathfinish_subpath pathstart_subpath reversepath_joinpaths by (metis (no_types, lifting)) (*Below F cannot be from 'a \ 'b because the dot product won't work. We have that g returns 'a and then F takes the output of g, so F should start from 'a Then we have to compute the dot product of the vector b with both the derivative of g, and F. Since the derivative of g returns the same type as g, accordingly F should return the same type as g, i.e. 'a. *) definition line_integral:: "('a::euclidean_space \ 'a::euclidean_space) \ (('a) set) \ (real \ 'a) \ real" where "line_integral F basis g \ integral {0 .. 1} (\x. \b\basis. (F(g x) \ b) * (vector_derivative g (at x within {0..1}) \ b))" definition line_integral_exists where "line_integral_exists F basis \ \ (\x. \b\basis. F(\ x) \ b * (vector_derivative \ (at x within {0..1}) \ b)) integrable_on {0..1}" lemma line_integral_on_pair_straight_path: fixes F::"('a::euclidean_space) \ 'a" and g :: "real \ real" and \ assumes gamma_const: "\x. \(x)\ i = a" and gamma_smooth: "\x \ {0 .. 1}. \ differentiable at x" shows "(line_integral F {i} \) = 0" "(line_integral_exists F {i} \)" proof (simp add: line_integral_def) have *: "F (\ x) \ i * (vector_derivative \ (at x within {0..1}) \ i) = 0" if "0 \ x \ x \ 1" for x proof - have "((\x. \(x)\ i) has_vector_derivative 0) (at x)" using vector_derivative_const_at[of "a" "x"] and gamma_const by auto then have "(vector_derivative \ (at x) \ i) = 0" using derivative_component_fun_component[ of "\" "x" "i"] and gamma_smooth and that by (simp add: vector_derivative_at) then have "(vector_derivative \ (at x within {0 .. 1}) \ i) = 0" using has_vector_derivative_at_within vector_derivative_at_within_ivl that by (metis atLeastAtMost_iff gamma_smooth vector_derivative_works zero_less_one) then show ?thesis by auto qed then have "((\x. F (\ x) \ i * (vector_derivative \ (at x within {0..1}) \ i)) has_integral 0) {0..1}" using has_integral_is_0[of "{0 .. 1}" "(\x. F (\ x) \ i * (vector_derivative \ (at x within {0..1}) \ i))"] by auto then have "((\x. F (\ x) \ i * (vector_derivative \ (at x within {0..1}) \ i)) integrable_on {0..1})" by auto then show "line_integral_exists F {i} \" by (auto simp add:line_integral_exists_def) show "integral {0..1} (\x. F (\ x) \ i * (vector_derivative \ (at x within {0..1}) \ i)) = 0" using * has_integral_is_0[of "{0 .. 1}" "(\x. F (\ x) \ i * (vector_derivative \ (at x within {0..1}) \ i))"] by auto qed lemma line_integral_on_pair_path_strong: fixes F::"('a::euclidean_space) \ ('a)" and g::"real \ 'a" and \::"(real \ 'a)" and i::'a assumes i_norm_1: "norm i = 1" and g_orthogonal_to_i: "\x. g(x) \ i = 0" and gamma_is_in_terms_of_i: "\ = (\x. f(x) *\<^sub>R i + g(f(x)))" and gamma_smooth: "\ piecewise_C1_differentiable_on {0 .. 1}" and g_continuous_on_f: "continuous_on (f ` {0..1}) g" and path_start_le_path_end: "(pathstart \) \ i \ (pathfinish \) \ i" and field_i_comp_cont: "continuous_on (path_image \) (\x. F x \ i)" shows "line_integral F {i} \ = integral (cbox ((pathstart \) \ i) ((pathfinish \) \ i)) (\f_var. (F (f_var *\<^sub>R i + g(f_var)) \ i))" "line_integral_exists F {i} \" proof (simp add: line_integral_def) obtain s where gamma_differentiable: "finite s" "(\x \ {0 .. 1} - s. \ differentiable at x)" using gamma_smooth by (auto simp add: C1_differentiable_on_eq piecewise_C1_differentiable_on_def) then have gamma_i_component_smooth: "\x \ {0 .. 1} - s. (\x. \ x \ i) differentiable at x" by auto have field_cont_on_path: "continuous_on ((\x. \ x \ i) ` {0..1}) (\f_var. F (f_var *\<^sub>R i + g f_var) \ i)" proof - have 0: "(\x. \ x \ i) = f" proof fix x show "\ x \ i = f x" using g_orthogonal_to_i i_norm_1 by (simp only: gamma_is_in_terms_of_i real_inner_class.inner_add_left g_orthogonal_to_i inner_scaleR_left inner_same_Basis norm_eq_1) qed show ?thesis unfolding 0 apply (rule continuous_on_compose2 [of _ "(\x. F(x) \ i)" "f ` { 0..1}" "(\x. x *\<^sub>R i + g x)"] field_i_comp_cont g_continuous_on_f field_i_comp_cont continuous_intros)+ by (auto simp add: gamma_is_in_terms_of_i path_image_def) qed have path_start_le_path_end': "\ 0 \ i \ \ 1 \ i" using path_start_le_path_end by (auto simp add: pathstart_def pathfinish_def) have gamm_cont: "continuous_on {0..1} (\a. \ a \ i)" apply(rule continuous_on_inner) using gamma_smooth apply (simp add: piecewise_C1_differentiable_on_def) using continuous_on_const by auto then obtain c d where cd: "c \ d" "(\a. \ a \ i) ` {0..1} = {c..d}" by (meson continuous_image_closed_interval zero_le_one) then have subset_cd: "(\a. \ a \ i) ` {0..1} \ {c..d}" by auto have field_cont_on_path_cd: "continuous_on {c..d} (\f_var. F (f_var *\<^sub>R i + g f_var) \ i)" using field_cont_on_path cd by auto have path_vector_deriv_line_integrals: "\x\{0..1} - s. ((\x. \ x \ i) has_vector_derivative vector_derivative (\x. \ x \ i) (at x)) (at x)" using gamma_i_component_smooth and derivative_component_fun_component and vector_derivative_works by blast then have "\x\{0..1} - s. ((\x. \ x \ i) has_vector_derivative vector_derivative (\x. \ x \ i) (at x within ({0..1}))) (at x within ({0..1}))" using has_vector_derivative_at_within vector_derivative_at_within_ivl by fastforce then have has_int:"((\x. vector_derivative (\x. \ x \ i) (at x within {0..1}) *\<^sub>R (F ((\ x \ i) *\<^sub>R i + g (\ x \ i)) \ i)) has_integral integral {\ 0 \ i..\ 1 \ i} (\f_var. F (f_var *\<^sub>R i + g f_var) \ i)) {0..1}" using has_integral_substitution_strong[OF gamma_differentiable(1) rel_simps(44) path_start_le_path_end' subset_cd field_cont_on_path_cd gamm_cont, of "(\x. vector_derivative (\x. \(x) \ i) (at x within ({0..1})))"] gamma_is_in_terms_of_i by (auto simp only: has_field_derivative_iff_has_vector_derivative) then have has_int':"((\x. (F(\(x)) \ i)*(vector_derivative (\x. \(x) \ i) (at x within ({0..1})))) has_integral integral {((pathstart \) \ i)..((pathfinish \) \ i)} (\f_var. F (f_var *\<^sub>R i + g f_var) \ i)) {0..1}" using gamma_is_in_terms_of_i i_norm_1 apply (auto simp add: pathstart_def pathfinish_def) apply (simp only: real_inner_class.inner_add_left inner_not_same_Basis g_orthogonal_to_i inner_scaleR_left norm_eq_1) by (auto simp add: algebra_simps) have substitute: "integral ({((pathstart \) \ i)..((pathfinish \) \ i)}) (\f_var. (F (f_var *\<^sub>R i + g(f_var)) \ i)) = integral ({0..1}) (\x. (F(\(x)) \ i)*(vector_derivative (\x. \(x) \ i) (at x within ({0..1}))))" using gamma_is_in_terms_of_i integral_unique[OF has_int] i_norm_1 apply (auto simp add: pathstart_def pathfinish_def) apply (simp only: real_inner_class.inner_add_left inner_not_same_Basis g_orthogonal_to_i inner_scaleR_left norm_eq_1) by (auto simp add: algebra_simps) have comp_in_eq_comp_out: "\x \ {0..1} - s. (vector_derivative (\x. \(x) \ i) (at x within {0..1})) = (vector_derivative \ (at x within {0..1})) \ i" proof fix x:: real assume ass:"x \ {0..1} -s" then have x_bounds:"x \ {0..1}" by auto have "\ differentiable at x" using ass gamma_differentiable by auto then have dotprod_in_is_out: "vector_derivative (\x. \(x) \ i) (at x) = (vector_derivative \ (at x)) \ i" using derivative_component_fun_component by force then have 0: "(vector_derivative \ (at x)) \ i = (vector_derivative \ (at x within {0..1})) \ i" proof - have "(\ has_vector_derivative vector_derivative \ (at x)) (at x)" using \\ differentiable at x\ vector_derivative_works by blast moreover have "0 \ x \ x \ 1" using x_bounds by presburger ultimately show ?thesis by (simp add: vector_derivative_at_within_ivl) qed have 1: "vector_derivative (\x. \(x) \ i) (at x) = vector_derivative (\x. \(x) \ i) (at x within {0..1})" using path_vector_deriv_line_integrals and vector_derivative_at_within_ivl and x_bounds by (metis ass atLeastAtMost_iff zero_less_one) show "vector_derivative (\x. \ x \ i) (at x within {0..1}) = vector_derivative \ (at x within {0..1}) \ i" using 0 and 1 and dotprod_in_is_out by auto qed show "integral {0..1} (\x. F (\ x) \ i * (vector_derivative \ (at x within {0..1}) \ i)) = integral {pathstart \ \ i..pathfinish \ \ i} (\f_var. F (f_var *\<^sub>R i + g f_var) \ i)" using substitute and comp_in_eq_comp_out and negligible_finite Henstock_Kurzweil_Integration.integral_spike [of "s" "{0..1}" "(\x. F (\ x) \ i * (vector_derivative \ (at x within {0..1}) \ i))" "(\x. F (\ x) \ i * vector_derivative (\x. \ x \ i) (at x within {0..1}))"] by (metis (no_types, lifting) gamma_differentiable(1)) have "((\x. F (\ x) \ i * (vector_derivative \ (at x within {0..1}) \ i)) has_integral integral {pathstart \ \ i..pathfinish \ \ i} (\f_var. F (f_var *\<^sub>R i + g f_var) \ i)) {0..1}" using has_int' and comp_in_eq_comp_out and negligible_finite Henstock_Kurzweil_Integration.has_integral_spike [of "s" "{0..1}" "(\x. F (\ x) \ i * (vector_derivative \ (at x within {0..1}) \ i))" "(\x. F (\ x) \ i * vector_derivative (\x. \ x \ i) (at x within {0..1}))" "integral {pathstart \ \ i..pathfinish \ \ i} (\f_var. F (f_var *\<^sub>R i + g f_var) \ i)"] by (metis (no_types, lifting) gamma_differentiable(1)) then have "(\x. F (\ x) \ i * (vector_derivative \ (at x within {0..1}) \ i)) integrable_on {0..1}" using integrable_on_def by auto then show "line_integral_exists F {i} \" by (auto simp add: line_integral_exists_def) qed lemma line_integral_on_pair_path: fixes F::"('a::euclidean_space) \ ('a)" and g::"real \ 'a" and \::"(real \ 'a)" and i::'a assumes i_norm_1: "norm i = 1" and g_orthogonal_to_i: "\x. g(x) \ i = 0" and gamma_is_in_terms_of_i: "\ = (\x. f(x) *\<^sub>R i + g(f(x)))" and gamma_smooth: "\ C1_differentiable_on {0 .. 1}" and g_continuous_on_f: "continuous_on (f ` {0..1}) g" and path_start_le_path_end: "(pathstart \) \ i \ (pathfinish \) \ i" and field_i_comp_cont: "continuous_on (path_image \) (\x. F x \ i)" shows "(line_integral F {i} \) = integral (cbox ((pathstart \) \ i) ((pathfinish \) \ i)) (\f_var. (F (f_var *\<^sub>R i + g(f_var)) \ i))" proof (simp add: line_integral_def) have gamma_differentiable: "\x \ {0 .. 1}. \ differentiable at x" using gamma_smooth C1_differentiable_on_eq by blast then have gamma_i_component_smooth: "\x \ {0 .. 1}. (\x. \ x \ i) differentiable at x" by auto have vec_deriv_i_comp_cont: "continuous_on {0..1} (\x. vector_derivative (\x. \ x \ i) (at x within {0..1}))" proof - have "continuous_on {0..1} (\x. vector_derivative (\x. \ x) (at x within {0..1}))" using gamma_smooth C1_differentiable_on_eq by (smt C1_differentiable_on_def atLeastAtMost_iff continuous_on_eq vector_derivative_at_within_ivl) then have deriv_comp_cont: "continuous_on {0..1} (\x. vector_derivative (\x. \ x) (at x within {0..1}) \ i)" by (simp add: continuous_intros) show ?thesis using derivative_component_fun_component_at_within[OF gamma_differentiable, of "i"] continuous_on_eq[OF deriv_comp_cont, of "(\x. vector_derivative (\x. \ x \ i) (at x within {0..1}))"] by fastforce qed have field_cont_on_path: "continuous_on ((\x. \ x \ i) ` {0..1}) (\f_var. F (f_var *\<^sub>R i + g f_var) \ i)" proof - have 0: "(\x. \ x \ i) = f" proof fix x show "\ x \ i = f x" using g_orthogonal_to_i i_norm_1 by (simp only: gamma_is_in_terms_of_i real_inner_class.inner_add_left g_orthogonal_to_i inner_scaleR_left inner_same_Basis norm_eq_1) qed show ?thesis unfolding 0 apply (rule continuous_on_compose2 [of _ "(\x. F(x) \ i)" "f ` { 0..1}" "(\x. x *\<^sub>R i + g x)"] field_i_comp_cont g_continuous_on_f field_i_comp_cont continuous_intros)+ by (auto simp add: gamma_is_in_terms_of_i path_image_def) qed have path_vector_deriv_line_integrals: "\x\{0..1}. ((\x. \ x \ i) has_vector_derivative vector_derivative (\x. \ x \ i) (at x)) (at x)" using gamma_i_component_smooth and derivative_component_fun_component and vector_derivative_works by blast then have "\x\{0..1}. ((\x. \ x \ i) has_vector_derivative vector_derivative (\x. \ x \ i) (at x within {0..1})) (at x within {0..1})" using has_vector_derivative_at_within vector_derivative_at_within_ivl by fastforce then have substitute: "integral (cbox ((pathstart \) \ i) ((pathfinish \) \ i)) (\f_var. (F (f_var *\<^sub>R i + g(f_var)) \ i)) = integral (cbox 0 1) (\x. (F(\(x)) \ i)*(vector_derivative (\x. \(x) \ i) (at x within {0..1})))" using gauge_integral_by_substitution [of "0" "1" "(\x. (\ x) \ i)" "(\x. vector_derivative (\x. \(x) \ i) (at x within {0..1}))" "(\f_var. (F (f_var *\<^sub>R i + g(f_var)) \ i))"] and path_start_le_path_end and vec_deriv_i_comp_cont and field_cont_on_path and gamma_is_in_terms_of_i i_norm_1 apply (auto simp add: pathstart_def pathfinish_def) apply (simp only: real_inner_class.inner_add_left inner_not_same_Basis g_orthogonal_to_i inner_scaleR_left norm_eq_1) by (auto) (*integration by substitution*) have comp_in_eq_comp_out: "\x \ {0..1}. (vector_derivative (\x. \(x) \ i) (at x within {0..1})) = (vector_derivative \ (at x within {0..1})) \ i" proof fix x:: real assume x_bounds: "x \ {0..1}" then have "\ differentiable at x" using gamma_differentiable by auto then have dotprod_in_is_out: "vector_derivative (\x. \(x) \ i) (at x) = (vector_derivative \ (at x)) \ i" using derivative_component_fun_component by force then have 0: "(vector_derivative \ (at x)) \ i = (vector_derivative \ (at x within {0..1})) \ i" using has_vector_derivative_at_within and x_bounds and vector_derivative_at_within_ivl by (smt atLeastAtMost_iff gamma_differentiable inner_commute vector_derivative_works) have 1: "vector_derivative (\x. \(x) \ i) (at x) = vector_derivative (\x. \(x) \ i) (at x within {0..1})" using path_vector_deriv_line_integrals and vector_derivative_at_within_ivl and x_bounds by fastforce show "vector_derivative (\x. \ x \ i) (at x within {0..1}) = vector_derivative \ (at x within {0..1}) \ i" using 0 and 1 and dotprod_in_is_out by auto qed show "integral {0..1} (\x. F (\ x) \ i * (vector_derivative \ (at x within {0..1}) \ i)) = integral {pathstart \ \ i..pathfinish \ \ i} (\f_var. F (f_var *\<^sub>R i + g f_var) \ i)" using substitute and comp_in_eq_comp_out and Henstock_Kurzweil_Integration.integral_cong [of "{0..1}" "(\x. F (\ x) \ i * (vector_derivative \ (at x within {0..1}) \ i))" "(\x. F (\ x) \ i * vector_derivative (\x. \ x \ i) (at x within {0..1}))"] by auto qed lemma content_box_cases: "content (box a b) = (if \i\Basis. a\i \ b\i then prod (\i. b\i - a\i) Basis else 0)" by (simp add: measure_lborel_box_eq inner_diff) lemma content_box_cbox: shows "content (box a b) = content (cbox a b)" by (simp add: content_box_cases content_cbox_cases) lemma content_eq_0: "content (box a b) = 0 \ (\i\Basis. b\i \ a\i)" by (auto simp: content_box_cases not_le intro: less_imp_le antisym eq_refl) lemma content_pos_lt_eq: "0 < content (cbox a (b::'a::euclidean_space)) \ (\i\Basis. a\i < b\i)" by (auto simp add: content_cbox_cases less_le prod_nonneg) lemma content_lt_nz: "0 < content (box a b) \ content (box a b) \ 0" using Paths.content_eq_0 zero_less_measure_iff by blast lemma content_subset: "cbox a b \ box c d \ content (cbox a b) \ content (box c d)" unfolding measure_def by (intro enn2real_mono emeasure_mono) (auto simp: emeasure_lborel_cbox_eq emeasure_lborel_box_eq) lemma sum_content_null: assumes "content (box a b) = 0" and "p tagged_division_of (box a b)" shows "sum (\(x,k). content k *\<^sub>R f x) p = (0::'a::real_normed_vector)" proof (rule sum.neutral, rule) fix y assume y: "y \ p" obtain x k where xk: "y = (x, k)" using surj_pair[of y] by blast note assm = tagged_division_ofD(3-4)[OF assms(2) y[unfolded xk]] from this(2) obtain c d where k: "k = cbox c d" by blast have "(\(x, k). content k *\<^sub>R f x) y = content k *\<^sub>R f x" unfolding xk by auto also have "\ = 0" using content_subset[OF assm(1)[unfolded k]] content_pos_le[of "cbox c d"] unfolding assms(1) k by auto finally show "(\(x, k). content k *\<^sub>R f x) y = 0" . qed lemma has_integral_null [intro]: "content(box a b) = 0 \ (f has_integral 0) (box a b)" by (simp add: content_box_cbox content_eq_0_interior) lemma line_integral_distrib: assumes "line_integral_exists f basis g1" "line_integral_exists f basis g2" "valid_path g1" "valid_path g2" shows "line_integral f basis (g1 +++ g2) = line_integral f basis g1 + line_integral f basis g2" "line_integral_exists f basis (g1 +++ g2)" proof - obtain s1 s2 where s1: "finite s1" "\x\{0..1} - s1. g1 differentiable at x" and s2: "finite s2" "\x\{0..1} - s2. g2 differentiable at x" using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) obtain i1 i2 where 1: "((\x. \b\basis. f (g1 x) \ b * (vector_derivative g1 (at x within {0..1}) \ b)) has_integral i1) {0..1}" and 2: "((\x. \b\basis. f (g2 x) \ b * (vector_derivative g2 (at x within {0..1}) \ b)) has_integral i2) {0..1}" using assms by (auto simp: line_integral_exists_def) have i1: "((\x. 2 * (\b\basis. f (g1 (2 * x)) \ b * (vector_derivative g1 (at (2 * x) within {0..1}) \ b))) has_integral i1) {0..1/2}" and i2: "((\x. 2 * (\b\basis. f (g2 (2 * x - 1)) \ b * (vector_derivative g2 (at ((2 * x) - 1) within {0..1}) \ b))) has_integral i2) {1/2..1}" using has_integral_affinity01 [OF 1, where m= 2 and c=0, THEN has_integral_cmul [where c=2]] has_integral_affinity01 [OF 2, where m= 2 and c="-1", THEN has_integral_cmul [where c=2]] by (simp_all only: image_affinity_atLeastAtMost_div_diff, simp_all add: scaleR_conv_of_real mult_ac) have g1: "\0 \ z; z*2 < 1; z*2 \ s1\ \ vector_derivative (\x. if x*2 \ 1 then g1 (2*x) else g2 (2*x - 1)) (at z within {0..1}) = 2 *\<^sub>R vector_derivative g1 (at (z*2) within {0..1})" for z proof - have i:"\0 \ z; z*2 < 1; z*2 \ s1\ \ vector_derivative (\x. if x * 2 \ 1 then g1 (2 * x) else g2 (2 * x - 1)) (at z within {0..1}) = 2 *\<^sub>R vector_derivative g1 (at (z * 2))" for z proof- have g1_at_z:"\0 \ z; z*2 < 1; z*2 \ s1\ \ ((\x. if x*2 \ 1 then g1 (2*x) else g2 (2*x - 1)) has_vector_derivative 2 *\<^sub>R vector_derivative g1 (at (z*2))) (at z)" for z apply (rule has_vector_derivative_transform_at [of "\z - 1/2\" _ "(\x. g1(2*x))"]) apply (simp_all add: dist_real_def abs_if split: if_split_asm) apply (rule vector_diff_chain_at [of "\x. 2*x" 2 _ g1, simplified o_def]) apply (simp add: has_vector_derivative_def has_derivative_def bounded_linear_mult_left) using s1 apply (auto simp: algebra_simps vector_derivative_works) done assume ass: "0 \ z" "z*2 < 1" "z*2 \ s1" then have z_ge: "z\ 1" by auto show "vector_derivative (\x. if x * 2 \ 1 then g1 (2 * x) else g2 (2 * x - 1)) (at z within {0..1}) = 2 *\<^sub>R vector_derivative g1 (at (z * 2))" using Derivative.vector_derivative_at_within_ivl[OF g1_at_z[OF ass] ass(1) z_ge] by auto qed assume ass: "0 \ z" "z*2 < 1" "z*2 \ s1" then have "(g1 has_vector_derivative ((vector_derivative g1 (at (z*2))))) (at (z*2))" using s1 by (auto simp: algebra_simps vector_derivative_works) then have ii: "(vector_derivative g1 (at (z*2) within {0..1})) = (vector_derivative g1 (at (z*2)))" using Derivative.vector_derivative_at_within_ivl ass by force show "vector_derivative (\x. if x * 2 \ 1 then g1 (2 * x) else g2 (2 * x - 1)) (at z within {0..1}) = 2 *\<^sub>R vector_derivative g1 (at (z * 2) within {0..1})" using i[OF ass] ii by auto qed have g2: "\1 < z*2; z \ 1; z*2 - 1 \ s2\ \ vector_derivative (\x. if x*2 \ 1 then g1 (2*x) else g2 (2*x - 1)) (at z within {0..1}) = 2 *\<^sub>R vector_derivative g2 (at (z*2 - 1) within {0..1})" for z proof - have i:"\1 < z*2; z \ 1; z*2 - 1 \ s2\ \ vector_derivative (\x. if x * 2 \ 1 then g1 (2 * x) else g2 (2 * x - 1)) (at z within {0..1}) = 2 *\<^sub>R vector_derivative g2 (at (z * 2 - 1))" for z proof- have g2_at_z:"\1 < z*2; z \ 1; z*2 - 1 \ s2\ \ ((\x. if x*2 \ 1 then g1 (2*x) else g2 (2*x - 1)) has_vector_derivative 2 *\<^sub>R vector_derivative g2 (at (z*2 - 1))) (at z)" for z apply (rule has_vector_derivative_transform_at [of "\z - 1/2\" _ "(\x. g2 (2*x - 1))"]) apply (simp_all add: dist_real_def abs_if split: if_split_asm) apply (rule vector_diff_chain_at [of "\x. 2*x - 1" 2 _ g2, simplified o_def]) apply (simp add: has_vector_derivative_def has_derivative_def bounded_linear_mult_left) using s2 apply (auto simp: algebra_simps vector_derivative_works) done assume ass: "1 < z*2" "z \ 1" "z*2 - 1 \ s2" then have z_le: "z\ 1" by auto have z_ge: "0 \ z" using ass by auto show "vector_derivative (\x. if x * 2 \ 1 then g1 (2 * x) else g2 (2 * x - 1)) (at z within {0..1}) = 2 *\<^sub>R vector_derivative g2 (at (z * 2 - 1))" using Derivative.vector_derivative_at_within_ivl[OF g2_at_z[OF ass] z_ge z_le] by auto qed assume ass: "1 < z*2" "z \ 1" "z*2 - 1 \ s2" then have "(g2 has_vector_derivative ((vector_derivative g2 (at (z*2 - 1))))) (at (z*2 - 1))" using s2 by (auto simp: algebra_simps vector_derivative_works) then have ii: "(vector_derivative g2 (at (z*2 - 1) within {0..1})) = (vector_derivative g2 (at (z*2 - 1)))" using Derivative.vector_derivative_at_within_ivl ass by force show "vector_derivative (\x. if x * 2 \ 1 then g1 (2 * x) else g2 (2 * x - 1)) (at z within {0..1}) = 2 *\<^sub>R vector_derivative g2 (at (z * 2 - 1) within {0..1})" using i[OF ass] ii by auto qed have lem1: "((\x. \b\basis. f ((g1+++g2) x) \ b * (vector_derivative (g1+++g2) (at x within {0..1}) \ b)) has_integral i1) {0..1/2}" apply (rule has_integral_spike_finite [OF _ _ i1, of "insert (1/2) ((*)2 -` s1)"]) using s1 apply (force intro: finite_vimageI [where h = "(*)2"] inj_onI) apply (clarsimp simp add: joinpaths_def scaleR_conv_of_real mult_ac g1) by (simp add: sum_distrib_left) moreover have lem2: "((\x. \b\basis. f ((g1+++g2) x) \ b * (vector_derivative (g1+++g2) (at x within {0..1}) \ b)) has_integral i2) {1/2..1}" apply (rule has_integral_spike_finite [OF _ _ i2, of "insert (1/2) ((\x. 2*x-1) -` s2)"]) using s2 apply (force intro: finite_vimageI [where h = "\x. 2*x-1"] inj_onI) apply (clarsimp simp add: joinpaths_def scaleR_conv_of_real mult_ac g2) by (simp add: sum_distrib_left) ultimately show "line_integral f basis (g1 +++ g2) = line_integral f basis g1 + line_integral f basis g2" apply (simp add: line_integral_def) apply (rule integral_unique [OF has_integral_combine [where c = "1/2"]]) using 1 2 integral_unique apply auto done show "line_integral_exists f basis (g1 +++ g2)" proof (simp add: line_integral_exists_def integrable_on_def) have "(1::real) \ 1 * 2 \ (0::real) \ 1 / 2" by simp then show "\r. ((\r. \a\basis. f ((g1 +++ g2) r) \ a * (vector_derivative (g1 +++ g2) (at r within {0..1}) \ a)) has_integral r) {0..1}" using has_integral_combine [where c = "1/2"] 1 2 divide_le_eq_numeral1(1) lem1 lem2 by blast qed qed lemma line_integral_exists_joinD1: assumes "line_integral_exists f basis (g1 +++ g2)" "valid_path g1" shows "line_integral_exists f basis g1" proof - obtain s1 where s1: "finite s1" "\x\{0..1} - s1. g1 differentiable at x" using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) have "(\x. \b\basis. f ((g1 +++ g2) (x/2)) \ b * (vector_derivative (g1 +++ g2) (at (x/2) within {0..1}) \ b)) integrable_on {0..1}" using assms apply (auto simp: line_integral_exists_def) apply (drule integrable_on_subcbox [where a=0 and b="1/2"]) apply (auto intro: integrable_affinity [of _ 0 "1/2::real" "1/2" 0, simplified]) done then have *:"(\x. \b\basis. ((f ((g1 +++ g2) (x/2)) \ b) / 2)* (vector_derivative (g1 +++ g2) (at (x/2) within {0..1}) \ b)) integrable_on {0..1}" by (auto simp: Groups_Big.sum_distrib_left dest: integrable_cmul [where c="1/2"] simp: scaleR_conv_of_real) have g1: "\0 \ z; z*2 < 1; z*2 \ s1\ \ vector_derivative (\x. if x*2 \ 1 then g1 (2*x) else g2 (2*x - 1)) (at z within {0..1}) = 2 *\<^sub>R vector_derivative g1 (at (z*2) within {0..1})" for z proof - have i:"\0 \ z; z*2 < 1; z*2 \ s1\ \ vector_derivative (\x. if x * 2 \ 1 then g1 (2 * x) else g2 (2 * x - 1)) (at z within {0..1}) = 2 *\<^sub>R vector_derivative g1 (at (z * 2))" for z proof- have g1_at_z:"\0 \ z; z*2 < 1; z*2 \ s1\ \ ((\x. if x*2 \ 1 then g1 (2*x) else g2 (2*x - 1)) has_vector_derivative 2 *\<^sub>R vector_derivative g1 (at (z*2))) (at z)" for z apply (rule has_vector_derivative_transform_at [of "\z - 1/2\" _ "(\x. g1(2*x))"]) apply (simp_all add: dist_real_def abs_if split: if_split_asm) apply (rule vector_diff_chain_at [of "\x. 2*x" 2 _ g1, simplified o_def]) apply (simp add: has_vector_derivative_def has_derivative_def bounded_linear_mult_left) using s1 apply (auto simp: algebra_simps vector_derivative_works) done assume ass: "0 \ z" "z*2 < 1" "z*2 \ s1" then have z_ge: "z\ 1" by auto show "vector_derivative (\x. if x * 2 \ 1 then g1 (2 * x) else g2 (2 * x - 1)) (at z within {0..1}) = 2 *\<^sub>R vector_derivative g1 (at (z * 2))" using Derivative.vector_derivative_at_within_ivl[OF g1_at_z[OF ass] ass(1) z_ge] by auto qed assume ass: "0 \ z" "z*2 < 1" "z*2 \ s1" then have "(g1 has_vector_derivative ((vector_derivative g1 (at (z*2))))) (at (z*2))" using s1 by (auto simp: algebra_simps vector_derivative_works) then have ii: "(vector_derivative g1 (at (z*2) within {0..1})) = (vector_derivative g1 (at (z*2)))" using Derivative.vector_derivative_at_within_ivl ass by force show "vector_derivative (\x. if x * 2 \ 1 then g1 (2 * x) else g2 (2 * x - 1)) (at z within {0..1}) = 2 *\<^sub>R vector_derivative g1 (at (z * 2) within {0..1})" using i[OF ass] ii by auto qed show ?thesis using s1 apply (auto simp: line_integral_exists_def) apply (rule integrable_spike_finite [of "{0,1} \ s1", OF _ _ *]) apply (auto simp: joinpaths_def scaleR_conv_of_real g1) done qed lemma line_integral_exists_joinD2: assumes "line_integral_exists f basis (g1 +++ g2)" "valid_path g2" shows "line_integral_exists f basis g2" proof - obtain s2 where s2: "finite s2" "\x\{0..1} - s2. g2 differentiable at x" using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) have "(\x. \b\basis. f ((g1 +++ g2) (x/2 + 1/2)) \ b * (vector_derivative (g1 +++ g2) (at (x/2 + 1/2) within {0..1}) \ b)) integrable_on {0..1}" using assms apply (auto simp: line_integral_exists_def) apply (drule integrable_on_subcbox [where a="1/2" and b=1], auto) apply (drule integrable_affinity [of _ "1/2::real" 1 "1/2" "1/2", simplified]) apply (simp add: image_affinity_atLeastAtMost_diff) done then have *:"(\x. \b\basis. ((f ((g1 +++ g2) (x/2 + 1/2)) \ b) / 2)* (vector_derivative (g1 +++ g2) (at (x/2 + 1/2) within {0..1}) \ b)) integrable_on {0..1}" by (auto simp: Groups_Big.sum_distrib_left dest: integrable_cmul [where c="1/2"] simp: scaleR_conv_of_real) have g2: "\1 < z*2; z \ 1; z*2 - 1 \ s2\ \ vector_derivative (\x. if x*2 \ 1 then g1 (2*x) else g2 (2*x - 1)) (at z within {0..1}) = 2 *\<^sub>R vector_derivative g2 (at (z*2 - 1) within {0..1})" for z proof - have i:"\1 < z*2; z \ 1; z*2 - 1 \ s2\ \ vector_derivative (\x. if x * 2 \ 1 then g1 (2 * x) else g2 (2 * x - 1)) (at z within {0..1}) = 2 *\<^sub>R vector_derivative g2 (at (z * 2 - 1))" for z proof- have g2_at_z:"\1 < z*2; z \ 1; z*2 - 1 \ s2\ \ ((\x. if x*2 \ 1 then g1 (2*x) else g2 (2*x - 1)) has_vector_derivative 2 *\<^sub>R vector_derivative g2 (at (z*2 - 1))) (at z)" for z apply (rule has_vector_derivative_transform_at [of "\z - 1/2\" _ "(\x. g2 (2*x - 1))"]) apply (simp_all add: dist_real_def abs_if split: if_split_asm) apply (rule vector_diff_chain_at [of "\x. 2*x - 1" 2 _ g2, simplified o_def]) apply (simp add: has_vector_derivative_def has_derivative_def bounded_linear_mult_left) using s2 apply (auto simp: algebra_simps vector_derivative_works) done assume ass: "1 < z*2" "z \ 1" "z*2 - 1 \ s2" then have z_le: "z\ 1" by auto have z_ge: "0 \ z" using ass by auto show "vector_derivative (\x. if x * 2 \ 1 then g1 (2 * x) else g2 (2 * x - 1)) (at z within {0..1}) = 2 *\<^sub>R vector_derivative g2 (at (z * 2 - 1))" using Derivative.vector_derivative_at_within_ivl[OF g2_at_z[OF ass] z_ge z_le] by auto qed assume ass: "1 < z*2" "z \ 1" "z*2 - 1 \ s2" then have "(g2 has_vector_derivative ((vector_derivative g2 (at (z*2 - 1))))) (at (z*2 - 1))" using s2 by (auto simp: algebra_simps vector_derivative_works) then have ii: "(vector_derivative g2 (at (z*2 - 1) within {0..1})) = (vector_derivative g2 (at (z*2 - 1)))" using Derivative.vector_derivative_at_within_ivl ass by force show "vector_derivative (\x. if x * 2 \ 1 then g1 (2 * x) else g2 (2 * x - 1)) (at z within {0..1}) = 2 *\<^sub>R vector_derivative g2 (at (z * 2 - 1) within {0..1})" using i[OF ass] ii by auto qed show ?thesis using s2 apply (auto simp: line_integral_exists_def) apply (rule integrable_spike_finite [of "{0,1} \ s2", OF _ _ *]) apply (auto simp: joinpaths_def scaleR_conv_of_real g2) done qed lemma has_line_integral_on_reverse_path: assumes g: "valid_path g" and int: "((\x. \b\basis. F (g x) \ b * (vector_derivative g (at x within {0..1}) \ b)) has_integral c){0..1}" shows "((\x. \b\basis. F ((reversepath g) x) \ b * (vector_derivative (reversepath g) (at x within {0..1}) \ b)) has_integral -c){0..1}" proof - { fix s x assume xs: "g C1_differentiable_on ({0..1} - s)" "x \ (-) 1 ` s" "0 \ x" "x \ 1" have "vector_derivative (\x. g (1 - x)) (at x within {0..1}) = - vector_derivative g (at (1 - x) within {0..1})" proof - obtain f' where f': "(g has_vector_derivative f') (at (1 - x))" using xs by (force simp: has_vector_derivative_def C1_differentiable_on_def) have "(g o (\x. 1 - x) has_vector_derivative -1 *\<^sub>R f') (at x)" apply (rule vector_diff_chain_within) apply (intro vector_diff_chain_within derivative_eq_intros | simp)+ apply (rule has_vector_derivative_at_within [OF f']) done then have mf': "((\x. g (1 - x)) has_vector_derivative -f') (at x)" by (simp add: o_def) show ?thesis using xs by (auto simp: vector_derivative_at_within_ivl [OF mf'] vector_derivative_at_within_ivl [OF f']) qed } note * = this obtain S where "continuous_on {0..1} g" "finite S" "g C1_differentiable_on {0..1} - S" using g by (auto simp: valid_path_def piecewise_C1_differentiable_on_def) then show ?thesis using has_integral_affinity01 [OF int, where m= "-1" and c=1] unfolding reversepath_def by (rule_tac S = "(\x. 1 - x) ` S" in has_integral_spike_finite) (auto simp: * has_integral_neg Groups_Big.sum_negf) qed lemma line_integral_on_reverse_path: assumes "valid_path \" "line_integral_exists F basis \" shows "line_integral F basis \ = - (line_integral F basis (reversepath \))" "line_integral_exists F basis (reversepath \)" proof - obtain i where 0: "((\x. \b\basis. F (\ x) \ b * (vector_derivative \ (at x within {0..1}) \ b)) has_integral i){0..1}" using assms unfolding integrable_on_def line_integral_exists_def by auto then have 1: "((\x. \b\basis. F ((reversepath \) x) \ b * (vector_derivative (reversepath \) (at x within {0..1}) \ b)) has_integral -i){0..1}" using has_line_integral_on_reverse_path assms by auto then have rev_line_integral:"line_integral F basis (reversepath \) = -i" using line_integral_def Henstock_Kurzweil_Integration.integral_unique by (metis (no_types)) have line_integral: "line_integral F basis \ = i" using line_integral_def 0 Henstock_Kurzweil_Integration.integral_unique by blast show "line_integral F basis \ = - (line_integral F basis (reversepath \))" using line_integral rev_line_integral by auto show "line_integral_exists F basis (reversepath \)" using 1 line_integral_exists_def by auto qed lemma line_integral_exists_on_degenerate_path: assumes "finite basis" shows "line_integral_exists F basis (\x. c)" proof- have every_component_integrable: "\b\basis. (\x. F ((\x. c) x) \ b * (vector_derivative (\x. c) (at x within {0..1}) \ b)) integrable_on {0..1}" proof fix b assume b_in_basis: "b \ basis" have cont_field_zero_one: "continuous_on {0..1} (\x. F ((\x. c) x) \ b)" using continuous_on_const by fastforce have cont_path_zero_one: "continuous_on {0..1} (\x. (vector_derivative (\x. c) (at x within {0..1})) \ b)" proof - have "((vector_derivative (\x. c) (at x within {0..1})) \ b) = 0" if "x \ {0..1}" for x proof - have "vector_derivative (\x. c) (at x within {0..1}) = 0" using that gamma_deriv_at_within[of "0" "1"] differentiable_const vector_derivative_const_at by fastforce then show "vector_derivative (\x. c) (at x within {0..1}) \ b = 0" by auto qed then show "continuous_on {0..1} (\x. (vector_derivative (\x. c) (at x within {0..1})) \ b)" using continuous_on_const[of "{0..1}" "0"] continuous_on_eq[of "{0..1}" "\x. 0" "(\x. (vector_derivative (\x. c) (at x within {0..1})) \ b)"] by auto qed show "(\x. F (c) \ b * (vector_derivative (\x. c) (at x within {0..1}) \ b)) integrable_on {0..1}" using cont_field_zero_one cont_path_zero_one continuous_on_mult integrable_continuous_real by blast qed have integrable_sum': "\t f s. finite t \ \a\t. f a integrable_on s \ (\x. \a\t. f a x) integrable_on s" using integrable_sum by metis have field_integrable_on_basis: "(\x. \b\basis. F (c) \ b * (vector_derivative (\x. c) (at x within {0..1}) \ b)) integrable_on {0..1}" using integrable_sum'[OF assms(1) every_component_integrable] by auto then show ?thesis using line_integral_exists_def by auto qed lemma degenerate_path_is_valid_path: "valid_path (\x. c)" by (auto simp add: valid_path_def piecewise_C1_differentiable_on_def continuous_on_const) lemma line_integral_degenerate_path: assumes "finite basis" shows "line_integral F basis (\x. c) = 0" proof (simp add: line_integral_def) have "((vector_derivative (\x. c) (at x within {0..1})) \ b) = 0" if "x \ {0..1}" for x b proof - have "vector_derivative (\x. c) (at x within {0..1}) = 0" using that gamma_deriv_at_within[of "0" "1"] differentiable_const vector_derivative_const_at by fastforce then show "vector_derivative (\x. c) (at x within {0..1}) \ b = 0" by auto qed then have 0: "\x. x \ {0..1} \ (\x. \b\basis. F c \ b * (vector_derivative (\x. c) (at x within {0..1}) \ b)) x = (\x. 0) x" by auto then show "integral {0..1} (\x. \b\basis. F c \ b * (vector_derivative (\x. c) (at x within {0..1}) \ b)) = 0" using integral_cong[of "{0..1}", OF 0] integral_0 by auto qed definition point_path where "point_path \ \ \c. \ = (\x. c)" lemma line_integral_point_path: assumes "point_path \" assumes "finite basis" shows "line_integral F basis \ = 0" using assms(1) point_path_def line_integral_degenerate_path[OF assms(2)] by force lemma line_integral_exists_point_path: assumes "finite basis" "point_path \" shows "line_integral_exists F basis \" using assms apply(simp add: point_path_def) using line_integral_exists_on_degenerate_path by auto lemma line_integral_exists_subpath: assumes f: "line_integral_exists f basis g" and g: "valid_path g" and uv: "u \ {0..1}" "v \ {0..1}" "u \ v" shows "(line_integral_exists f basis (subpath u v g))" proof (cases "v=u") case tr: True have zero: "(\b\basis. f (g u) \ b * (vector_derivative (\x. g u) (at x within {0..1}) \ b)) = 0" if "x \ {0..1}" for x::real proof - have "(vector_derivative (\x. g u) (at x within {0..1})) = 0" using Deriv.has_vector_derivative_const that Derivative.vector_derivative_at_within_ivl by fastforce then show "(\b\basis. f (g u) \ b * (vector_derivative (\x. g u) (at x within {0..1}) \ b)) = 0" by auto qed then have "((\x. \b\basis. f (g u) \ b * (vector_derivative (\x. g u) (at x within {0..1}) \ b)) has_integral 0) {0..1}" by (meson has_integral_is_0) then show ?thesis using f tr by (auto simp add: line_integral_def line_integral_exists_def subpath_def) next case False obtain s where s: "\x. x \ {0..1} - s \ g differentiable at x" and fs: "finite s" using g unfolding piecewise_C1_differentiable_on_def C1_differentiable_on_eq valid_path_def by blast have *: "((\x. \b\basis. f (g ((v - u) * x + u)) \ b * (vector_derivative g (at ((v - u) * x + u) within {0..1}) \ b)) has_integral (1 / (v - u)) * integral {u..v} (\x. \b\basis. f (g (x)) \ b * (vector_derivative g (at x within {0..1}) \ b))) {0..1}" using f uv apply (simp add: line_integral_exists_def subpath_def) apply (drule integrable_on_subcbox [where a=u and b=v, simplified]) apply (simp_all add: has_integral_integral) apply (drule has_integral_affinity [where m="v-u" and c=u, simplified]) apply (simp_all add: False image_affinity_atLeastAtMost_div_diff scaleR_conv_of_real) apply (simp add: divide_simps False) done have vd:"\x. x \ {0..1} \ x \ (\t. (v-u) *\<^sub>R t + u) -` s \ vector_derivative (\x. g ((v-u) * x + u)) (at x within {0..1}) = (v-u) *\<^sub>R vector_derivative g (at ((v-u) * x + u) within {0..1})" using test2[OF s fs uv] by auto have arg:"\x. (\n\basis. (v - u) * (f (g ((v - u) * x + u)) \ n) * (vector_derivative g (at ((v - u) * x + u) within {0..1}) \ n)) = (\b\basis. f (g ((v - u) * x + u)) \ b * (v - u) * (vector_derivative g (at ((v - u) * x + u) within {0..1}) \ b))" by (simp add: mult.commute) have"((\x. \b\basis. f (g ((v - u) * x + u)) \ b * (vector_derivative (\x. g ((v - u) * x + u)) (at x within {0..1}) \ b)) has_integral (integral {u..v} (\x. \b\basis. f (g (x)) \ b * (vector_derivative g (at x within {0..1}) \ b)))) {0..1}" apply (cut_tac Henstock_Kurzweil_Integration.has_integral_mult_right [OF *, where c = "v-u"]) using fs assms apply (simp add: False subpath_def line_integral_exists_def) apply (rule_tac S = "(\t. ((v-u) *\<^sub>R t + u)) -` s" in has_integral_spike_finite) apply (auto simp: inj_on_def False vd finite_vimageI scaleR_conv_of_real Groups_Big.sum_distrib_left mult.assoc[symmetric] arg) done then show "(line_integral_exists f basis (subpath u v g))" by(auto simp add: line_integral_exists_def subpath_def integrable_on_def) qed (*This should have everything that has to do with onecubes*) type_synonym path = "real \ (real * real)" type_synonym one_cube = "(real \ (real * real))" type_synonym one_chain = "(int * path) set" type_synonym two_cube = "(real * real) \ (real * real)" type_synonym two_chain = "two_cube set" definition one_chain_line_integral :: "((real * real) \ (real * real)) => ((real*real) set) \ one_chain \ real" where "one_chain_line_integral F b C \ (\(k,g)\C. k * (line_integral F b g))" definition boundary_chain where "boundary_chain s \ (\(k, \) \ s. k = 1 \ k = -1)" fun coeff_cube_to_path::"(int * one_cube) \ path" where "coeff_cube_to_path (k, \) = (if k = 1 then \ else (reversepath \))" fun rec_join :: "(int*path) list \ path" where "rec_join [] = (\x.0)" | "rec_join [oneC] = coeff_cube_to_path oneC" | "rec_join (oneC # xs) = coeff_cube_to_path oneC +++ (rec_join xs)" fun valid_chain_list where "valid_chain_list [] = True" | "valid_chain_list [oneC] = True" | "valid_chain_list (oneC#l) = (pathfinish (coeff_cube_to_path (oneC)) = pathstart (rec_join l) \ valid_chain_list l)" lemma joined_is_valid: assumes boundary_chain: "boundary_chain (set l)" and valid_path: "\k \. (k, \) \ set l \ valid_path \" and valid_chain_list_ass: "valid_chain_list l" shows "valid_path (rec_join l)" using assms proof(induction l) case Nil then show ?case using C1_differentiable_imp_piecewise[OF C1_differentiable_on_const[of "0" "{0..1}"]] by (auto simp add: valid_path_def) next case (Cons a l) have *: "valid_path (rec_join ((k::int, \) # l))" if "boundary_chain (set (l))" "(\k' \'. (k', \') \ set l \ valid_path \')" "valid_chain_list l" "valid_path (rec_join l) " "(\k' \'. (k', \') \ set ((k, \) # l) \ valid_path \')" "valid_chain_list ((k, \) # l)" "boundary_chain (set ((k,\) # l))" for k \ l proof (cases "l = []") case True with that show "valid_path (rec_join ((k, \) # l))" by auto next case False then obtain h l' where l_nempty: "l = h#l'" by (meson rec_join.elims) then show "valid_path (rec_join ((k, \) # l))" proof (simp, intro conjI impI) assume k_eq_1: "k = 1" have 0:"valid_path \" using that by auto have 1:"pathfinish \ = pathstart (rec_join (h#l'))" using that(6) k_eq_1 l_nempty by auto show "valid_path (\ +++ rec_join (h#l'))" using 0 1 valid_path_join that(4) l_nempty by auto next assume "k \ 1" then have k_eq_neg_1: "k = -1" using that(7) by (auto simp add: boundary_chain_def) have "valid_path \" using that by auto then have 0: "valid_path (reversepath \)" - using Cauchy_Integral_Theorem.valid_path_imp_reverse + using valid_path_imp_reverse by auto have 1: "pathfinish (reversepath \) = pathstart (rec_join (h#l'))" using that(6) k_eq_neg_1 l_nempty by auto show "valid_path ((reversepath \) +++ rec_join (h#l'))" - using 0 1 Cauchy_Integral_Theorem.valid_path_join that(4) l_nempty by blast + using 0 1 valid_path_join that(4) l_nempty by blast qed qed have "\ps. valid_chain_list ps \ (\i f p psa. ps = (i, f) # p # psa \ ((i = 1 \ pathfinish f \ pathstart (rec_join (p # psa)) \ i \ 1 \ pathfinish (reversepath f) \ pathstart (rec_join (p # psa))) \ \ valid_chain_list (p # psa)))" by (smt coeff_cube_to_path.elims valid_chain_list.elims(3)) moreover have "boundary_chain (set l)" by (meson Cons.prems(1) boundary_chain_def set_subset_Cons subset_eq) ultimately show ?case using * Cons by (metis (no_types) list.set_intros(2) prod.collapse valid_chain_list.simps(3)) qed lemma pathstart_rec_join_1: "pathstart (rec_join ((1, \) # l)) = pathstart \" proof (cases "l = []") case True then show "pathstart (rec_join ((1, \) # l)) = pathstart \" by simp next case False then obtain h l' where "l = h#l'" by (meson rec_join.elims) then show "pathstart (rec_join ((1, \) # l)) = pathstart \" by simp qed lemma pathstart_rec_join_2: "pathstart (rec_join ((-1, \) # l)) = pathstart (reversepath \)" proof cases assume "l = []" then show "pathstart (rec_join ((- 1, \) # l)) = pathstart (reversepath \)" by simp next assume "l \ [] " then obtain h l' where "l = h#l'" by (meson rec_join.elims) then show "pathstart (rec_join ((- 1, \) # l)) = pathstart (reversepath \)" by simp qed lemma pathstart_rec_join: "pathstart (rec_join ((1, \) # l)) = pathstart \" "pathstart (rec_join ((-1, \) # l)) = pathstart (reversepath \)" using pathstart_rec_join_1 pathstart_rec_join_2 by auto lemma line_integral_exists_on_rec_join: assumes boundary_chain: "boundary_chain (set l)" and valid_chain_list: "valid_chain_list l" and valid_path: "\k \. (k, \) \ set l \ valid_path \" and line_integral_exists: "\(k, \) \ set l. line_integral_exists F basis \" shows "line_integral_exists F basis (rec_join l)" using assms proof (induction l) case Nil then show ?case proof (simp add: line_integral_exists_def) have "\x. (vector_derivative (\x. 0) (at x)) = 0" using Derivative.vector_derivative_const_at by auto then have "\x. ((\x. 0) has_vector_derivative 0) (at x)" using Derivative.vector_derivative_const_at by auto then have "\x. ((\x. 0) has_vector_derivative 0) (at x within {0..1})" using Derivative.vector_derivative_const_at by auto then have 0: "\x\{0..1}. (vector_derivative (\x. 0) (at x within{0..1})) = 0" by (simp add: gamma_deriv_at_within) have "(\x\{0..1}. (\b\basis. F 0 \ b * (vector_derivative (\x. 0) (at x within {0..1}) \ b)) = 0)" by (simp add: 0) then have " ((\x. \b\basis. F 0 \ b * (vector_derivative (\x. 0) (at x within {0..1}) \ b)) has_integral 0) {0..1}" by (meson has_integral_is_0) then show "(\x. \b\basis. F 0 \ b * (vector_derivative (\x. 0) (at x within {0..1}) \ b)) integrable_on {0..1}" by auto qed next case (Cons a l) obtain k \ where aeq: "a = (k,\)" by force show ?case unfolding aeq proof cases assume l_empty: "l = []" then show "line_integral_exists F basis (rec_join ((k, \) # l))" using Cons.prems aeq line_integral_on_reverse_path(2) by fastforce next assume "l \ []" then obtain h l' where l_nempty: "l = h#l'" by (meson rec_join.elims) show "line_integral_exists F basis (rec_join ((k, \) # l))" proof (auto simp add: l_nempty) assume k_eq_1: "k = 1" have 0: "line_integral_exists F basis \" using Cons.prems(4) aeq by auto have 1: "line_integral_exists F basis (rec_join l)" by (metis (mono_tags) Cons boundary_chain_def list.set_intros(2) valid_chain_list.elims(3) valid_chain_list.simps(3)) have 2: "valid_path \" using Cons aeq by auto have 3:"valid_path (rec_join l)" by (metis (no_types) Cons.prems boundary_chain_def joined_is_valid l_nempty set_subset_Cons subsetCE valid_chain_list.simps(3)) show "line_integral_exists F basis (\ +++ rec_join (h#l'))" using line_integral_distrib(2)[OF 0 1 2 3] assms l_nempty by auto next assume "k \ 1" then have k_eq_neg_1: "k = -1" using Cons aeq by (simp add: boundary_chain_def) have gamma_valid: "valid_path \" using Cons aeq by auto then have 2: "valid_path (reversepath \)" - using Cauchy_Integral_Theorem.valid_path_imp_reverse by auto + using valid_path_imp_reverse by auto have "line_integral_exists F basis \" using Cons aeq by auto then have 0: "line_integral_exists F basis (reversepath \)" using line_integral_on_reverse_path(2) gamma_valid by auto have 1: "line_integral_exists F basis (rec_join l)" using Cons aeq by (metis (mono_tags) boundary_chain_def insert_iff list.set(2) list.set_intros(2) valid_chain_list.elims(3) valid_chain_list.simps(3)) have 3:"valid_path (rec_join l)" by (metis (no_types) Cons.prems(1) Cons.prems(2) Cons.prems(3) boundary_chain_def joined_is_valid l_nempty set_subset_Cons subsetCE valid_chain_list.simps(3)) show "line_integral_exists F basis ((reversepath \) +++ rec_join (h#l'))" using line_integral_distrib(2)[OF 0 1 2 3] assms l_nempty by auto qed qed qed lemma line_integral_exists_rec_join_cons: assumes "line_integral_exists F basis (rec_join ((1,\) # l))" "(\k' \'. (k', \') \ set ((1,\) # l) \ valid_path \')" "finite basis" shows "line_integral_exists F basis (\ +++ (rec_join l))" proof cases assume l_empty: "l = []" show "line_integral_exists F basis (\ +++ rec_join l)" using assms(2) line_integral_distrib(2)[OF assms(1) line_integral_exists_on_degenerate_path[OF assms(3)], of "0"] using degenerate_path_is_valid_path by (fastforce simp add: l_empty) next assume "l \ []" then obtain h l' where "l = h#l'" by (meson rec_join.elims) then show "line_integral_exists F basis (\ +++ rec_join l)" using assms by auto qed lemma line_integral_exists_rec_join_cons_2: assumes "line_integral_exists F basis (rec_join ((-1,\) # l))" "(\k' \'. (k', \') \ set ((1,\) # l) \ valid_path \')" "finite basis" shows "line_integral_exists F basis ((reversepath \) +++ (rec_join l))" proof cases assume l_empty: "l = []" show "line_integral_exists F basis ((reversepath \) +++ rec_join l)" using assms(2) line_integral_distrib(2)[OF assms(1) line_integral_exists_on_degenerate_path[OF assms(3)], of "0"] using degenerate_path_is_valid_path by (auto simp add: l_empty) next assume "l \ []" then obtain h l' where "l = h#l'" by (meson rec_join.elims) with assms show "line_integral_exists F basis ((reversepath \) +++ rec_join l)" using assms by auto qed lemma line_integral_exists_on_rec_join': assumes boundary_chain: "boundary_chain (set l)" and valid_chain_list: "valid_chain_list l" and valid_path: "\k \. (k, \) \ set l \ valid_path \" and line_integral_exists: "line_integral_exists F basis (rec_join l)" and finite_basis: "finite basis" shows "\(k, \) \ set l. line_integral_exists F basis \" using assms proof (induction l) case Nil show ?case by (simp add: line_integral_exists_def) next case ass: (Cons a l) obtain k \ where k_gamma:"a = (k,\)" by fastforce show ?case apply (auto simp add: k_gamma) proof - show "line_integral_exists F basis \" proof(cases "k = 1") assume k_eq_1: "k = 1" have 0: "line_integral_exists F basis (\ +++ (rec_join l))" using line_integral_exists_rec_join_cons k_eq_1 k_gamma ass(4) ass(5) ass(6) by auto have 2: "valid_path \" using ass k_gamma by auto show "line_integral_exists F basis \" using line_integral_exists_joinD1[OF 0 2] by auto next assume "k \ 1" then have k_eq_neg_1: "k = -1" using ass k_gamma by (simp add: boundary_chain_def) have 0: "line_integral_exists F basis ((reversepath \) +++ (rec_join l))" using line_integral_exists_rec_join_cons_2[OF ] k_eq_neg_1 k_gamma ass(4) ass(5) ass(6) by fastforce have gamma_valid: "valid_path \" using ass k_gamma by auto then have 2: "valid_path (reversepath \)" - using Cauchy_Integral_Theorem.valid_path_imp_reverse by auto + using valid_path_imp_reverse by auto have "line_integral_exists F basis (reversepath \)" using line_integral_exists_joinD1[OF 0 2] by auto then show "line_integral_exists F basis (\)" using line_integral_on_reverse_path(2)[OF 2] reversepath_reversepath by auto qed next have 0:"boundary_chain (set l)" using ass(2) by (auto simp add: boundary_chain_def) have 1:"valid_chain_list l" using ass(3) apply (auto simp add: k_gamma) by (metis valid_chain_list.elims(3) valid_chain_list.simps(3)) have 2:"(\k \. (k, \) \ set (l) \ valid_path \)" using ass(4) by auto have 3: "valid_path (rec_join l)" using joined_is_valid[OF 0] 1 2 by auto have 4: "line_integral_exists F basis (rec_join l)" proof(cases "k = 1") assume k_eq_1: "k = 1" have 0: "line_integral_exists F basis (\ +++ (rec_join l))" using line_integral_exists_rec_join_cons k_eq_1 k_gamma ass(4) ass(5) ass(6) by auto show "line_integral_exists F basis (rec_join l)" using line_integral_exists_joinD2[OF 0 3] by auto next assume "k \ 1" then have k_eq_neg_1: "k = -1" using ass k_gamma by (simp add: boundary_chain_def) have 0: "line_integral_exists F basis ((reversepath \) +++ (rec_join l))" using line_integral_exists_rec_join_cons_2[OF ] k_eq_neg_1 k_gamma ass(4) ass(5) ass(6) by fastforce show "line_integral_exists F basis (rec_join l)" using line_integral_exists_joinD2[OF 0 3] by auto qed show "\a b. (a, b) \ set l \ line_integral_exists F basis b" using 0 1 2 3 4 ass(1)[OF 0 1 2] ass(6) by fastforce qed qed inductive chain_subdiv_path where I: "chain_subdiv_path \ (set l)" if "distinct l" "rec_join l = \" "valid_chain_list l" lemma valid_path_equiv_valid_chain_list: assumes path_eq_chain: "chain_subdiv_path \ one_chain" and "boundary_chain one_chain" "\(k, \) \ one_chain. valid_path \" shows "valid_path \" proof - obtain l where l_props: "set l = one_chain" "distinct l" "rec_join l = \" "valid_chain_list l" using chain_subdiv_path.cases path_eq_chain by force show "valid_path \" using joined_is_valid assms l_props by blast qed lemma line_integral_rec_join_cons: assumes "line_integral_exists F basis \" "line_integral_exists F basis (rec_join ((l)))" "(\k' \'. (k', \') \ set ((1,\) # l) \ valid_path \')" "finite basis" shows "line_integral F basis (rec_join ((1,\) # l)) = line_integral F basis (\ +++ (rec_join l))" proof cases assume l_empty: "l = []" show "line_integral F basis (rec_join ((1,\) # l)) = line_integral F basis (\ +++ (rec_join l))" using assms line_integral_distrib(1)[OF assms(1) line_integral_exists_on_degenerate_path[OF assms(4)], of "0"] apply (auto simp add: l_empty) using degenerate_path_is_valid_path line_integral_degenerate_path by fastforce next assume "l \ []" then obtain h l' where l_nempty: "l = h#l'" by (meson rec_join.elims) show "line_integral F basis (rec_join ((1,\) # l)) = line_integral F basis (\ +++ (rec_join l))" using assms by (auto simp add: l_nempty) qed lemma line_integral_rec_join_cons_2: assumes "line_integral_exists F basis \" "line_integral_exists F basis (rec_join ((l)))" "(\k' \'. (k', \') \ set ((-1,\) # l) \ valid_path \')" "finite basis" shows "line_integral F basis (rec_join ((-1,\) # l)) = line_integral F basis ((reversepath \) +++ (rec_join l))" proof cases assume l_empty: "l = []" have 0: "line_integral_exists F basis (reversepath \)" using assms line_integral_on_reverse_path(2) by fastforce have 1: "valid_path (reversepath \)" using assms by fastforce show "line_integral F basis (rec_join ((-1,\) # l)) = line_integral F basis ((reversepath \) +++ (rec_join l))" using assms line_integral_distrib(1)[OF 0 line_integral_exists_on_degenerate_path[OF assms(4)], of "0"] apply (auto simp add: l_empty) using degenerate_path_is_valid_path line_integral_degenerate_path by fastforce next assume "l \ []" then obtain h l' where l_nempty: "l = h#l'" by (meson rec_join.elims) show "line_integral F basis (rec_join ((-1,\) # l)) = line_integral F basis ((reversepath \) +++ (rec_join l))" using assms by (auto simp add: l_nempty) qed lemma one_chain_line_integral_rec_join: assumes l_props: "set l = one_chain" "distinct l" "valid_chain_list l" and boundary_chain: "boundary_chain one_chain" and line_integral_exists: "\(k::int, \) \ one_chain. line_integral_exists F basis \" and valid_path: "\(k::int, \) \ one_chain. valid_path \" and finite_basis: "finite basis" shows "line_integral F basis (rec_join l) = one_chain_line_integral F basis one_chain" proof - have 0: "sum_list (map (\((k::int), \). (k::int) * (line_integral F basis \)) l) = one_chain_line_integral F basis one_chain" unfolding one_chain_line_integral_def using l_props Groups_List.comm_monoid_add_class.sum.distinct_set_conv_list[OF l_props(2), of "(\(k, \). (k::int) * (line_integral F basis \))"] by auto have "valid_chain_list l \ boundary_chain (set l) \ (\(k::int, \) \ set l. line_integral_exists F basis \) \ (\(k::int, \) \ set l. valid_path \) \ line_integral F basis (rec_join l) = sum_list (map (\(k::int, \). k * (line_integral F basis \)) l)" proof (induction l) case Nil show ?case unfolding line_integral_def boundary_chain_def apply (auto) proof have "\x. (vector_derivative (\x. 0) (at x)) = 0" using Derivative.vector_derivative_const_at by auto then have "\x. ((\x. 0) has_vector_derivative 0) (at x)" using Derivative.vector_derivative_const_at by auto then have "\x. ((\x. 0) has_vector_derivative 0) (at x within {0..1})" using Derivative.vector_derivative_const_at by auto then have 0: "\x\{0..1}. (vector_derivative (\x. 0) (at x within{0..1})) = 0" by (metis (no_types) box_real(2) vector_derivative_within_cbox zero_less_one) have "(\x\{0..1}. (\b\basis. F 0 \ b * (vector_derivative (\x. 0) (at x within {0..1}) \ b)) = 0)" by (simp add: 0) then show "((\x. \b\basis. F 0 \ b * (vector_derivative (\x. 0) (at x within {0..1}) \ b)) has_integral 0) {0..1}" by (meson has_integral_is_0) qed next case ass: (Cons a l) obtain k::"int" and \::"one_cube" where props: "a = (k,\)" proof let ?k2 = "fst a" let ?\2 = "snd a" show "a = (?k2, ?\2)" by auto qed have "line_integral_exists F basis (rec_join (a # l))" using line_integral_exists_on_rec_join[OF ass(3) ass(2)] ass(5) ass(4) by blast have "boundary_chain (set l)" by (meson ass.prems(2) boundary_chain_def list.set_intros(2)) have val_l: "\f i. (i,f) \ set l \ valid_path f" using ass.prems(4) by fastforce have vcl_l: "valid_chain_list l" by (metis (no_types) ass.prems(1) valid_chain_list.elims(3) valid_chain_list.simps(3)) have line_integral_exists_on_joined: "line_integral_exists F basis (rec_join l)" by (metis \boundary_chain (set l)\ \line_integral_exists F basis (rec_join (a # l))\ emptyE val_l vcl_l joined_is_valid line_integral_exists_joinD2 line_integral_exists_on_rec_join list.set(1) neq_Nil_conv rec_join.simps(3)) have "valid_path (rec_join (a # l))" using joined_is_valid ass(5) ass(3) ass(2) by blast then have joined_is_valid: "valid_path (rec_join l)" using \boundary_chain (set l)\ val_l vcl_l joined_is_valid by blast show ?case proof (clarsimp, cases) assume k_eq_1: "(k::int) = 1" have line_integral_exists_on_gamma: "line_integral_exists F basis \" using ass props by auto have gamma_is_valid: "valid_path \" using ass props by auto have line_int_rw: "line_integral F basis (rec_join ((k, \) # l)) = line_integral F basis (\ +++ rec_join l)" proof - have gam_int: "line_integral_exists F basis \" using ass props by auto have rec_join_int: "line_integral_exists F basis (rec_join l)" using line_integral_exists_on_rec_join using line_integral_exists_on_joined by blast show ?thesis using line_integral_rec_join_cons[OF gam_int rec_join_int] ass k_eq_1 finite_basis props by force qed show "line_integral F basis (rec_join (a # l)) = (case a of (x, \) \ real_of_int x * line_integral F basis \) + (\(x, \)\l. real_of_int x * line_integral F basis \)" apply (simp add: props line_int_rw) using line_integral_distrib[OF line_integral_exists_on_gamma line_integral_exists_on_joined gamma_is_valid joined_is_valid] ass k_eq_1 vcl_l by (auto simp: boundary_chain_def props) next assume "k \ 1" then have k_eq_neg_1: "k = -1" using ass props by (auto simp add: boundary_chain_def) have line_integral_exists_on_gamma: "line_integral_exists F basis (reversepath \)" using line_integral_on_reverse_path ass props by auto have gamma_is_valid: "valid_path (reversepath \)" - using Cauchy_Integral_Theorem.valid_path_imp_reverse ass props by auto + using valid_path_imp_reverse ass props by auto have line_int_rw: "line_integral F basis (rec_join ((k, \) # l)) = line_integral F basis ((reversepath \) +++ rec_join l)" proof - have gam_int: "line_integral_exists F basis \" using ass props by auto have rec_join_int: "line_integral_exists F basis (rec_join l)" using line_integral_exists_on_rec_join using line_integral_exists_on_joined by blast show ?thesis using line_integral_rec_join_cons_2[OF gam_int rec_join_int] using ass k_eq_neg_1 using finite_basis props by blast qed show "line_integral F basis (rec_join (a # l)) = (case a of (x, \) \ real_of_int x * line_integral F basis \) + (\(x, \)\l. real_of_int x * line_integral F basis \)" apply (simp add: props line_int_rw) using line_integral_distrib[OF line_integral_exists_on_gamma line_integral_exists_on_joined gamma_is_valid joined_is_valid] props ass line_integral_on_reverse_path(1)[of "\" "F" "basis"] k_eq_neg_1 using \boundary_chain (set l)\ vcl_l by auto qed qed then have 1:"line_integral F basis (rec_join l) = sum_list (map (\(k::int, \). k * (line_integral F basis \)) l)" using l_props assms by auto then show ?thesis using 0 1 by auto qed lemma line_integral_on_path_eq_line_integral_on_equiv_chain: assumes path_eq_chain: "chain_subdiv_path \ one_chain" and boundary_chain: "boundary_chain one_chain" and line_integral_exists: "\(k::int, \) \ one_chain. line_integral_exists F basis \" and valid_path: "\(k::int, \) \ one_chain. valid_path \" and finite_basis: "finite basis" shows "one_chain_line_integral F basis one_chain = line_integral F basis \" "line_integral_exists F basis \" "valid_path \" proof - obtain l where l_props: "set l = one_chain" "distinct l" "rec_join l = \" "valid_chain_list l" using chain_subdiv_path.cases path_eq_chain by force show "line_integral_exists F basis \" using line_integral_exists_on_rec_join assms l_props by blast show "valid_path \" using joined_is_valid assms l_props by blast have "line_integral F basis (rec_join l) = one_chain_line_integral F basis one_chain" using one_chain_line_integral_rec_join l_props assms by auto then show "one_chain_line_integral F basis one_chain = line_integral F basis \" using l_props by auto qed lemma line_integral_on_path_eq_line_integral_on_equiv_chain': assumes path_eq_chain: "chain_subdiv_path \ one_chain" and boundary_chain: "boundary_chain one_chain" and line_integral_exists: "line_integral_exists F basis \" and valid_path: "\(k, \) \ one_chain. valid_path \" and finite_basis: "finite basis" shows "one_chain_line_integral F basis one_chain = line_integral F basis \" "\(k, \) \ one_chain. line_integral_exists F basis \" proof - obtain l where l_props: "set l = one_chain" "distinct l" "rec_join l = \" "valid_chain_list l" using chain_subdiv_path.cases path_eq_chain by force show 0: "\(k, \) \ one_chain. line_integral_exists F basis \" using line_integral_exists_on_rec_join' assms l_props by blast show "one_chain_line_integral F basis one_chain = line_integral F basis \" using line_integral_on_path_eq_line_integral_on_equiv_chain(1)[OF assms(1) assms(2) 0 assms(4) assms(5)] by auto qed definition chain_subdiv_chain where "chain_subdiv_chain one_chain1 subdiv \ \f. (\(f ` one_chain1)) = subdiv \ (\c\one_chain1. chain_subdiv_path (coeff_cube_to_path c) (f c)) \ pairwise (\ p p'. f p \ f p' = {}) one_chain1 \ (\x \ one_chain1. finite (f x))" lemma chain_subdiv_chain_character: shows "chain_subdiv_chain one_chain1 subdiv \ (\f. \(f ` one_chain1) = subdiv \ (\(k, \)\one_chain1. if k = 1 then chain_subdiv_path \ (f (k, \)) else chain_subdiv_path (reversepath \) (f (k, \))) \ (\p\one_chain1. \p'\one_chain1. p \ p' \ f p \ f p' = {}) \ (\x\one_chain1. finite (f x)))" unfolding chain_subdiv_chain_def by (safe; intro exI conjI iffI; fastforce simp add: pairwise_def) lemma chain_subdiv_chain_imp_finite_subdiv: assumes "finite one_chain1" "chain_subdiv_chain one_chain1 subdiv" shows "finite subdiv" using assms by(auto simp add: chain_subdiv_chain_def) lemma valid_subdiv_imp_valid_one_chain: assumes chain1_eq_chain2: "chain_subdiv_chain one_chain1 subdiv" and boundary_chain1: "boundary_chain one_chain1" and boundary_chain2: "boundary_chain subdiv" and valid_path: "\(k, \) \ subdiv. valid_path \" shows "\(k, \) \ one_chain1. valid_path \" proof - obtain f where f_props: "((\(f ` one_chain1)) = subdiv)" "(\(k,\)\one_chain1. if k = 1 then chain_subdiv_path \ (f(k,\)) else chain_subdiv_path (reversepath \) (f(k,\)))" "(\p\one_chain1. \p'\one_chain1. p \ p' \ f p \ f p' = {})" using chain1_eq_chain2 chain_subdiv_chain_character by auto have "\ k \. (k,\) \ one_chain1\ valid_path \" proof- fix k \ assume ass: "(k,\) \ one_chain1" show "valid_path \" proof (cases "k = 1") assume k_eq_1: "k = 1" then have i:"chain_subdiv_path \ (f(k,\))" using f_props(2) ass by auto have ii:"boundary_chain (f(k,\))" using f_props(1) ass assms apply (simp add: boundary_chain_def) by blast have "iv": " \(k, \)\f (k, \). valid_path \" using f_props(1) ass assms by blast show ?thesis using valid_path_equiv_valid_chain_list[OF i ii iv] by auto next assume "k \ 1" then have k_eq_neg1: "k = -1" using ass boundary_chain1 by (auto simp add: boundary_chain_def) then have i:"chain_subdiv_path (reversepath \) (f(k,\))" using f_props(2) ass using \k \ 1\ by auto have ii:"boundary_chain (f(k,\))" using f_props(1) ass assms by (auto simp add: boundary_chain_def) have "iv": " \(k, \)\f (k, \). valid_path \" using f_props(1) ass assms by blast have "valid_path (reversepath \)" using valid_path_equiv_valid_chain_list[OF i ii iv] by auto then show ?thesis - using reversepath_reversepath Cauchy_Integral_Theorem.valid_path_imp_reverse + using reversepath_reversepath valid_path_imp_reverse by force qed qed then show valid_path1: "\(k, \) \ one_chain1. valid_path \" by auto qed lemma one_chain_line_integral_eq_line_integral_on_sudivision: assumes chain1_eq_chain2: "chain_subdiv_chain one_chain1 subdiv" and boundary_chain1: "boundary_chain one_chain1" and boundary_chain2: "boundary_chain subdiv" and line_integral_exists_on_chain2: "\(k, \) \ subdiv. line_integral_exists F basis \" and valid_path: "\(k, \) \ subdiv. valid_path \" and finite_chain1: "finite one_chain1" and finite_basis: "finite basis" shows "one_chain_line_integral F basis one_chain1 = one_chain_line_integral F basis subdiv" "\(k, \) \ one_chain1. line_integral_exists F basis \" proof - obtain f where f_props: "((\(f ` one_chain1)) = subdiv)" "(\(k,\)\one_chain1. if k = 1 then chain_subdiv_path \ (f(k,\)) else chain_subdiv_path (reversepath \) (f(k,\)))" "(\p\one_chain1. \p'\one_chain1. p \ p' \ f p \ f p' = {})" "(\x \ one_chain1. finite (f x))" using chain1_eq_chain2 chain_subdiv_chain_character by (auto simp add: pairwise_def chain_subdiv_chain_def) then have 0: "one_chain_line_integral F basis subdiv = one_chain_line_integral F basis (\(f ` one_chain1))" by auto have finite_chain2: "finite subdiv" using finite_chain1 f_props(1) f_props(4) apply (simp add: image_def) using f_props(1) by auto have "\ k \. (k,\) \ one_chain1\ line_integral_exists F basis \" proof- fix k \ assume ass: "(k,\) \ one_chain1" show "line_integral_exists F basis \" proof (cases "k = 1") assume k_eq_1: "k = 1" then have i:"chain_subdiv_path \ (f(k,\))" using f_props(2) ass by auto have ii:"boundary_chain (f(k,\))" using f_props(1) ass assms by (auto simp add: boundary_chain_def) have iii:"\(k, \)\f (k, \). line_integral_exists F basis \" using f_props(1) ass assms by blast have "iv": " \(k, \)\f (k, \). valid_path \" using f_props(1) ass assms by blast show ?thesis using line_integral_on_path_eq_line_integral_on_equiv_chain(2)[OF i ii iii iv finite_basis] by auto next assume "k \ 1" then have k_eq_neg1: "k = -1" using ass boundary_chain1 by (auto simp add: boundary_chain_def) then have i:"chain_subdiv_path (reversepath \) (f(k,\))" using f_props(2) ass by auto have ii:"boundary_chain (f(k,\))" using f_props(1) ass assms by (auto simp add: boundary_chain_def) have iii:"\(k, \)\f (k, \). line_integral_exists F basis \" using f_props(1) ass assms by blast have "iv": " \(k, \)\f (k, \). valid_path \" using f_props(1) ass assms by blast have x:"line_integral_exists F basis (reversepath \)" using line_integral_on_path_eq_line_integral_on_equiv_chain(2)[OF i ii iii iv finite_basis] by auto have "valid_path (reversepath \)" using line_integral_on_path_eq_line_integral_on_equiv_chain(3)[OF i ii iii iv finite_basis] by auto then show ?thesis using line_integral_on_reverse_path(2) reversepath_reversepath x by fastforce qed qed then show line_integral_exists_on_chain1: "\(k, \) \ one_chain1. line_integral_exists F basis \" by auto have 1: "one_chain_line_integral F basis (\(f ` one_chain1)) = one_chain_line_integral F basis one_chain1" proof - have 0:"one_chain_line_integral F basis (\(f ` one_chain1)) = (\one_chain \ (f ` one_chain1). one_chain_line_integral F basis one_chain)" proof - have finite: "\chain \ (f ` one_chain1). finite chain" using f_props(1) finite_chain2 by (meson Sup_upper finite_subset) have disj: " \A\f ` one_chain1. \B\f ` one_chain1. A \ B \ A \ B = {}" by (metis (no_types, hide_lams) f_props(3) image_iff) show "one_chain_line_integral F basis (\(f ` one_chain1)) = (\one_chain \ (f ` one_chain1). one_chain_line_integral F basis one_chain)" using Groups_Big.comm_monoid_add_class.sum.Union_disjoint[OF finite disj] one_chain_line_integral_def by auto qed have 1:"(\one_chain \ (f ` one_chain1). one_chain_line_integral F basis one_chain) = one_chain_line_integral F basis one_chain1" proof - have "(\one_chain \ (f ` one_chain1). one_chain_line_integral F basis one_chain) = (\(k,\)\one_chain1. k*(line_integral F basis \))" proof - have i:"(\one_chain \ (f ` (one_chain1 - {p. f p = {}})). one_chain_line_integral F basis one_chain) = (\(k,\)\one_chain1 - {p. f p = {}}. k*(line_integral F basis \))" proof - have "inj_on f (one_chain1 - {p. f p = {}})" unfolding inj_on_def using f_props(3) by blast then have 0: "(\one_chain \ (f ` (one_chain1 - {p. f p = {}})). one_chain_line_integral F basis one_chain) = (\(k, \) \ (one_chain1 - {p. f p = {}}). one_chain_line_integral F basis (f (k, \)))" using Groups_Big.comm_monoid_add_class.sum.reindex by auto have "\ k \. (k, \) \ (one_chain1 - {p. f p = {}}) \ one_chain_line_integral F basis (f(k, \)) = k* (line_integral F basis \)" proof- fix k \ assume ass: "(k, \) \ (one_chain1 - {p. f p = {}})" have bchain: "boundary_chain (f(k,\))" using f_props(1) boundary_chain2 ass by (auto simp add: boundary_chain_def) have wexist: "\(k, \)\(f(k,\)). line_integral_exists F basis \" using f_props(1) line_integral_exists_on_chain2 ass by blast have vpath: " \(k, \)\(f(k, \)). valid_path \" using f_props(1) assms ass by blast show "one_chain_line_integral F basis (f (k, \)) = k * line_integral F basis \" proof(cases "k = 1") assume k_eq_1: "k = 1" have "one_chain_line_integral F basis (f (k, \)) = line_integral F basis \" using f_props(2) k_eq_1 line_integral_on_path_eq_line_integral_on_equiv_chain bchain wexist vpath ass finite_basis by auto then show "one_chain_line_integral F basis (f (k, \)) = k * line_integral F basis \" using k_eq_1 by auto next assume "k \ 1" then have k_eq_neg1: "k = -1" using ass boundary_chain1 by (auto simp add: boundary_chain_def) have "one_chain_line_integral F basis (f (k, \)) = line_integral F basis (reversepath \)" using f_props(2) k_eq_neg1 line_integral_on_path_eq_line_integral_on_equiv_chain bchain wexist vpath ass finite_basis by auto then have "one_chain_line_integral F basis (f (k, \)) = - (line_integral F basis \)" using line_integral_on_reverse_path(1) ass line_integral_exists_on_chain1 valid_subdiv_imp_valid_one_chain[OF chain1_eq_chain2 boundary_chain1 boundary_chain2 valid_path] by force then show "one_chain_line_integral F basis (f (k, \)) = k * line_integral F basis \" using k_eq_neg1 by auto qed qed then have "(\(k, \) \ (one_chain1 - {p. f p = {}}). one_chain_line_integral F basis (f (k, \))) = (\(k, \) \ (one_chain1 - {p. f p = {}}). k* (line_integral F basis \))" by (auto intro!: Finite_Cartesian_Product.sum_cong_aux) then show "(\one_chain \ (f ` (one_chain1 - {p. f p = {}})). one_chain_line_integral F basis one_chain) = (\(k, \) \ (one_chain1 - {p. f p = {}}). k* (line_integral F basis \))" using 0 by auto qed have "\ (k::int) \. (k, \) \ one_chain1 \ (f (k, \) = {}) \ (k, \) \ {(k',\'). k'* (line_integral F basis \') = 0}" proof- fix k::"int" fix \::"one_cube" assume ass:"(k, \) \ one_chain1" "(f (k, \) = {})" then have zero_line_integral:"one_chain_line_integral F basis (f (k, \)) = 0" using one_chain_line_integral_def by auto have bchain: "boundary_chain (f(k,\))" using f_props(1) boundary_chain2 ass by (auto simp add: boundary_chain_def) have wexist: "\(k, \)\(f(k,\)). line_integral_exists F basis \" using f_props(1) line_integral_exists_on_chain2 ass by blast have vpath: " \(k, \)\(f(k, \)). valid_path \" using f_props(1) assms ass by blast have "one_chain_line_integral F basis (f (k, \)) = k * line_integral F basis \" proof(cases "k = 1") assume k_eq_1: "k = 1" have "one_chain_line_integral F basis (f (k, \)) = line_integral F basis \" using f_props(2) k_eq_1 line_integral_on_path_eq_line_integral_on_equiv_chain bchain wexist vpath ass finite_basis by auto then show "one_chain_line_integral F basis (f (k, \)) = k * line_integral F basis \" using k_eq_1 by auto next assume "k \ 1" then have k_eq_neg1: "k = -1" using ass boundary_chain1 by (auto simp add: boundary_chain_def) have "one_chain_line_integral F basis (f (k, \)) = line_integral F basis (reversepath \)" using f_props(2) k_eq_neg1 line_integral_on_path_eq_line_integral_on_equiv_chain bchain wexist vpath ass finite_basis by auto then have "one_chain_line_integral F basis (f (k, \)) = - (line_integral F basis \)" using line_integral_on_reverse_path(1) ass line_integral_exists_on_chain1 valid_subdiv_imp_valid_one_chain[OF chain1_eq_chain2 boundary_chain1 boundary_chain2 valid_path] by force then show "one_chain_line_integral F basis (f (k::int, \)) = k * line_integral F basis \" using k_eq_neg1 by auto qed then show "(k, \) \ {(k'::int, \'). k' * line_integral F basis \' = 0}" using zero_line_integral by auto qed then have ii:"(\one_chain \ (f ` (one_chain1 - {p. f p = {}})). one_chain_line_integral F basis one_chain) = (\one_chain \ (f ` (one_chain1)). one_chain_line_integral F basis one_chain)" proof - have "\one_chain. one_chain \ (f ` (one_chain1)) - (f ` (one_chain1 - {p. f p = {}})) \ one_chain_line_integral F basis one_chain = 0" proof - fix one_chain assume "one_chain \ (f ` (one_chain1)) - (f ` (one_chain1 - {p. f p = {}}))" then show "one_chain_line_integral F basis one_chain = 0" by (auto simp add: one_chain_line_integral_def) qed then have 0:"(\one_chain \ f ` (one_chain1) - (f ` (one_chain1 - {p. f p = {}})). one_chain_line_integral F basis one_chain) = 0" using comm_monoid_add_class.sum.neutral by auto then have "(\one_chain \ f ` (one_chain1). one_chain_line_integral F basis one_chain) - (\one_chain \ (f ` (one_chain1 - {p. f p = {}})). one_chain_line_integral F basis one_chain) = 0" proof - have finte: "finite (f ` one_chain1)" using finite_chain1 by auto show ?thesis using Groups_Big.sum_diff[OF finte, of " (f ` (one_chain1 - {p. f p = {}}))" "one_chain_line_integral F basis"] 0 by auto qed then show "(\one_chain \ (f ` (one_chain1 - {p. f p = {}})). one_chain_line_integral F basis one_chain) = (\one_chain \ (f ` (one_chain1)). one_chain_line_integral F basis one_chain)" by auto qed have "\ (k::int) \. (k, \) \ one_chain1 \ (f (k, \) = {}) \ f (k, \) \ {chain. one_chain_line_integral F basis chain = 0}" proof- fix k::"int" fix \::"one_cube" assume ass:"(k, \) \ one_chain1" "(f (k, \) = {})" then have "one_chain_line_integral F basis (f (k, \)) = 0" using one_chain_line_integral_def by auto then show "f (k, \) \ {p'. (one_chain_line_integral F basis p' = 0)}" by auto qed then have iii:"(\(k::int,\)\one_chain1 - {p. f p = {}}. k*(line_integral F basis \)) = (\(k::int,\)\one_chain1. k*(line_integral F basis \))" proof- have "\ k \. (k,\)\one_chain1 - (one_chain1 - {p. f p = {}}) \ k* (line_integral F basis \) = 0" proof- fix k \ assume ass: "(k,\)\one_chain1 - (one_chain1 - {p. f p = {}})" then have "f(k, \) = {}" by auto then have "one_chain_line_integral F basis (f(k, \)) = 0" by (auto simp add: one_chain_line_integral_def) then have zero_line_integral:"one_chain_line_integral F basis (f (k, \)) = 0" using one_chain_line_integral_def by auto have bchain: "boundary_chain (f(k,\))" using f_props(1) boundary_chain2 ass by (auto simp add: boundary_chain_def) have wexist: "\(k, \)\(f(k,\)). line_integral_exists F basis \" using f_props(1) line_integral_exists_on_chain2 ass by blast have vpath: " \(k, \)\(f(k, \)). valid_path \" using f_props(1) assms ass by blast have "one_chain_line_integral F basis (f (k, \)) = k * line_integral F basis \" proof(cases "k = 1") assume k_eq_1: "k = 1" have "one_chain_line_integral F basis (f (k, \)) = line_integral F basis \" using f_props(2) k_eq_1 line_integral_on_path_eq_line_integral_on_equiv_chain bchain wexist vpath ass finite_basis by auto then show "one_chain_line_integral F basis (f (k, \)) = k * line_integral F basis \" using k_eq_1 by auto next assume "k \ 1" then have k_eq_neg1: "k = -1" using ass boundary_chain1 by (auto simp add: boundary_chain_def) have "one_chain_line_integral F basis (f (k, \)) = line_integral F basis (reversepath \)" using f_props(2) k_eq_neg1 line_integral_on_path_eq_line_integral_on_equiv_chain bchain wexist vpath ass finite_basis by auto then have "one_chain_line_integral F basis (f (k, \)) = - (line_integral F basis \)" using line_integral_on_reverse_path(1) ass line_integral_exists_on_chain1 valid_subdiv_imp_valid_one_chain[OF chain1_eq_chain2 boundary_chain1 boundary_chain2 valid_path] by force then show "one_chain_line_integral F basis (f (k, \)) = k * line_integral F basis \" using k_eq_neg1 by auto qed then show "k* (line_integral F basis \) = 0" using zero_line_integral by auto qed then have "\(k::int,\)\one_chain1 - (one_chain1 - {p. f p = {}}). k* (line_integral F basis \) = 0" by auto then have "(\(k::int,\)\one_chain1 - (one_chain1 - {p. f p = {}}). k*(line_integral F basis \)) = 0" using Groups_Big.comm_monoid_add_class.sum.neutral [of "one_chain1 - (one_chain1 - {p. f p = {}})" "(\(k::int,\). k* (line_integral F basis \))"] by (simp add: split_beta) then have "(\(k::int,\)\one_chain1. k*(line_integral F basis \)) - (\(k::int,\)\ (one_chain1 - {p. f p = {}}). k*(line_integral F basis \)) = 0" using Groups_Big.sum_diff[OF finite_chain1] by (metis (no_types) Diff_subset \(\(k, \)\one_chain1 - (one_chain1 - {p. f p = {}}). k * line_integral F basis \) = 0\ \\f B. B \ one_chain1 \ sum f (one_chain1 - B) = sum f one_chain1 - sum f B\) then show ?thesis by auto qed show ?thesis using i ii iii by auto qed then show ?thesis using one_chain_line_integral_def by auto qed show ?thesis using 0 1 by auto qed show "one_chain_line_integral F basis one_chain1 = one_chain_line_integral F basis subdiv" using 0 1 by auto qed lemma one_chain_line_integral_eq_line_integral_on_sudivision': assumes chain1_eq_chain2: "chain_subdiv_chain one_chain1 subdiv" and boundary_chain1: "boundary_chain one_chain1" and boundary_chain2: "boundary_chain subdiv" and line_integral_exists_on_chain1: "\(k, \) \ one_chain1. line_integral_exists F basis \" and valid_path: "\(k, \) \ subdiv. valid_path \" and finite_chain1: "finite one_chain1" and finite_basis: "finite basis" shows "one_chain_line_integral F basis one_chain1 = one_chain_line_integral F basis subdiv" "\(k, \) \ subdiv. line_integral_exists F basis \" proof - obtain f where f_props: "((\(f ` one_chain1)) = subdiv)" "(\(k,\)\one_chain1. if k = 1 then chain_subdiv_path \ (f(k,\)) else chain_subdiv_path (reversepath \) (f(k,\)))" "(\p\one_chain1. \p'\one_chain1. p \ p' \ f p \ f p' = {})" "(\x \ one_chain1. finite (f x))" using chain1_eq_chain2 chain_subdiv_chain_character by (auto simp add: pairwise_def chain_subdiv_chain_def) have finite_chain2: "finite subdiv" using finite_chain1 f_props(1) f_props(4) by blast have "\k \. (k, \) \ subdiv \ line_integral_exists F basis \" proof - fix k \ assume ass: "(k, \) \ subdiv" then obtain k' \' where kp_gammap: "(k',\') \ one_chain1" "(k,\) \ f(k',\')" using f_props(1) by fastforce show "line_integral_exists F basis \" proof (cases "k' = 1") assume k_eq_1: "k' = 1" then have i:"chain_subdiv_path \' (f(k',\'))" using f_props(2) kp_gammap ass by auto have ii:"boundary_chain (f(k',\'))" using f_props(1) ass assms kp_gammap by (meson UN_I boundary_chain_def) have iii:"line_integral_exists F basis \'" using assms kp_gammap by blast have "iv": " \(k, \)\f (k', \'). valid_path \" using f_props(1) ass assms kp_gammap by blast show ?thesis using line_integral_on_path_eq_line_integral_on_equiv_chain'(2)[OF i ii iii iv finite_basis] kp_gammap by auto next assume "k' \ 1" then have k_eq_neg1: "k' = -1" using boundary_chain1 kp_gammap by (auto simp add: boundary_chain_def) then have i:"chain_subdiv_path (reversepath \') (f(k',\'))" using f_props(2) kp_gammap by auto have ii:"boundary_chain (f(k',\'))" using f_props(1) assms kp_gammap by (meson UN_I boundary_chain_def) have iii: " \(k, \)\f (k', \'). valid_path \" using f_props(1) ass assms kp_gammap by blast have iv: "valid_path (reversepath \')" using valid_path_equiv_valid_chain_list[OF i ii iii] by force have "line_integral_exists F basis \'" using assms kp_gammap by blast then have x: "line_integral_exists F basis (reversepath \')" using iv line_integral_on_reverse_path(2) valid_path_reversepath by fastforce show ?thesis using line_integral_on_path_eq_line_integral_on_equiv_chain'(2)[OF i ii x iii finite_basis] kp_gammap by auto qed qed then show "\(k, \)\subdiv. line_integral_exists F basis \" by auto then show "one_chain_line_integral F basis one_chain1 = one_chain_line_integral F basis subdiv" using one_chain_line_integral_eq_line_integral_on_sudivision(1) assms by auto qed lemma line_integral_sum_gen: assumes finite_basis: "finite basis" and line_integral_exists: "line_integral_exists F basis1 \" "line_integral_exists F basis2 \" and basis_partition: "basis1 \ basis2 = basis" "basis1 \ basis2 = {}" shows "line_integral F basis \ = (line_integral F basis1 \) + (line_integral F basis2 \)" "line_integral_exists F basis \" apply (simp add: line_integral_def) proof - have 0: "integral {0..1} (\x. (\b\basis1. F (\ x) \ b * (vector_derivative \ (at x within {0..1}) \ b)) + (\b\basis2. F (\ x) \ b * (vector_derivative \ (at x within {0..1}) \ b))) = integral {0..1} (\x. \b\basis1. F (\ x) \ b * (vector_derivative \ (at x within {0..1}) \ b)) + integral {0..1} (\x. \b\basis2. F (\ x) \ b * (vector_derivative \ (at x within {0..1}) \ b))" using Henstock_Kurzweil_Integration.integral_add line_integral_exists by (auto simp add: line_integral_exists_def) have 1: "integral {0..1} (\x. \b\basis. F (\ x) \ b * (vector_derivative \ (at x within {0..1}) \ b)) = integral {0..1} (\x. (\b\basis1. F (\ x) \ b * (vector_derivative \ (at x within {0..1}) \ b)) + (\b\basis2. F (\ x) \ b * (vector_derivative \ (at x within {0..1}) \ b)))" by (metis (mono_tags, lifting) basis_partition finite_Un finite_basis sum.union_disjoint) show "integral {0..1} (\x. \b\basis. F (\ x) \ b * (vector_derivative \ (at x within {0..1}) \ b)) = integral {0..1} (\x. \b\basis1. F (\ x) \ b * (vector_derivative \ (at x within {0..1}) \ b)) + integral {0..1} (\x. \b\basis2. F (\ x) \ b * (vector_derivative \ (at x within {0..1}) \ b))" using 0 1 by auto have 2: "((\x. (\b\basis1. F (\ x) \ b * (vector_derivative \ (at x within {0..1}) \ b)) + (\b\basis2. F (\ x) \ b * (vector_derivative \ (at x within {0..1}) \ b))) has_integral integral {0..1} (\x. \b\basis1. F (\ x) \ b * (vector_derivative \ (at x within {0..1}) \ b)) + integral {0..1} (\x. \b\basis2. F (\ x) \ b * (vector_derivative \ (at x within {0..1}) \ b))) {0..1}" using Henstock_Kurzweil_Integration.has_integral_add line_integral_exists has_integral_integral apply (auto simp add: line_integral_exists_def) by blast have 3: "(\x. \b\basis. F (\ x) \ b * (vector_derivative \ (at x within {0..1}) \ b)) = (\x. (\b\basis1. F (\ x) \ b * (vector_derivative \ (at x within {0..1}) \ b)) + (\b\basis2. F (\ x) \ b * (vector_derivative \ (at x within {0..1}) \ b)))" by (metis (mono_tags, lifting) basis_partition finite_Un finite_basis sum.union_disjoint) show "line_integral_exists F basis \" apply(auto simp add: line_integral_exists_def has_integral_integral) using 2 3 using has_integral_integrable_integral by fastforce qed definition common_boundary_sudivision_exists where "common_boundary_sudivision_exists one_chain1 one_chain2 \ \subdiv. chain_subdiv_chain one_chain1 subdiv \ chain_subdiv_chain one_chain2 subdiv \ (\(k, \) \ subdiv. valid_path \)\ boundary_chain subdiv" lemma common_boundary_sudivision_commutative: "(common_boundary_sudivision_exists one_chain1 one_chain2) = (common_boundary_sudivision_exists one_chain2 one_chain1)" apply (simp add: common_boundary_sudivision_exists_def) by blast lemma common_subdivision_imp_eq_line_integral: assumes "(common_boundary_sudivision_exists one_chain1 one_chain2)" "boundary_chain one_chain1" "boundary_chain one_chain2" "\(k, \)\one_chain1. line_integral_exists F basis \" "finite one_chain1" "finite one_chain2" "finite basis" shows "one_chain_line_integral F basis one_chain1 = one_chain_line_integral F basis one_chain2" "\(k, \)\one_chain2. line_integral_exists F basis \" proof - obtain subdiv where subdiv_props: "chain_subdiv_chain one_chain1 subdiv" "chain_subdiv_chain one_chain2 subdiv" "(\(k, \) \ subdiv. valid_path \)" "(boundary_chain subdiv)" using assms by (auto simp add: common_boundary_sudivision_exists_def) have i: "\(k, \)\subdiv. line_integral_exists F basis \" using one_chain_line_integral_eq_line_integral_on_sudivision'(2)[OF subdiv_props(1) assms(2) subdiv_props(4) assms(4) subdiv_props(3) assms(5) assms(7)] by auto show "one_chain_line_integral F basis one_chain1 = one_chain_line_integral F basis one_chain2" using one_chain_line_integral_eq_line_integral_on_sudivision'(1)[OF subdiv_props(1) assms(2) subdiv_props(4) assms(4) subdiv_props(3) assms(5) assms(7)] one_chain_line_integral_eq_line_integral_on_sudivision(1)[OF subdiv_props(2) assms(3) subdiv_props(4) i subdiv_props(3) assms(6) assms(7)] by auto show "\(k, \)\one_chain2. line_integral_exists F basis \" using one_chain_line_integral_eq_line_integral_on_sudivision(2)[OF subdiv_props(2) assms(3) subdiv_props(4) i subdiv_props(3) assms(6) assms(7)] by auto qed definition common_sudiv_exists where "common_sudiv_exists one_chain1 one_chain2 \ \subdiv ps1 ps2. chain_subdiv_chain (one_chain1 - ps1) subdiv \ chain_subdiv_chain (one_chain2 - ps2) subdiv \ (\(k, \) \ subdiv. valid_path \) \ (boundary_chain subdiv) \ (\(k, \) \ ps1. point_path \) \ (\(k, \) \ ps2. point_path \)" lemma common_sudiv_exists_comm: shows "common_sudiv_exists C1 C2 = common_sudiv_exists C2 C1" by (auto simp add: common_sudiv_exists_def) lemma line_integral_degenerate_chain: assumes "(\(k, \) \ chain. point_path \)" assumes "finite basis" shows "one_chain_line_integral F basis chain = 0" proof (simp add: one_chain_line_integral_def) have "\(k,g)\chain. line_integral F basis g = 0" using assms line_integral_point_path by blast then have "\(k,g)\chain. real_of_int k * line_integral F basis g = 0" by auto then have "\p. p \ chain \ (case p of (i, f) \ real_of_int i * line_integral F basis f) = 0" by fastforce then show "(\x\chain. case x of (k, g) \ real_of_int k * line_integral F basis g) = 0" by simp qed lemma gen_common_subdiv_imp_common_subdiv: shows "(common_sudiv_exists one_chain1 one_chain2) = (\ps1 ps2. (common_boundary_sudivision_exists (one_chain1 - ps1) (one_chain2 - ps2)) \ (\(k, \)\ps1. point_path \) \ (\(k, \)\ps2. point_path \))" by (auto simp add: common_sudiv_exists_def common_boundary_sudivision_exists_def) lemma common_subdiv_imp_gen_common_subdiv: assumes "(common_boundary_sudivision_exists one_chain1 one_chain2)" shows "(common_sudiv_exists one_chain1 one_chain2)" using assms apply (auto simp add: common_sudiv_exists_def common_boundary_sudivision_exists_def) by (metis Diff_empty all_not_in_conv) lemma one_chain_line_integral_point_paths: assumes "finite one_chain" assumes "finite basis" assumes "(\(k, \)\ps. point_path \)" shows "one_chain_line_integral F basis (one_chain - ps) = one_chain_line_integral F basis (one_chain)" proof- have 0:"(\x\ps. case x of (k, g) \ (real_of_int k * line_integral F basis g) = 0)" using line_integral_point_path assms by force show "one_chain_line_integral F basis (one_chain - ps) = one_chain_line_integral F basis one_chain" unfolding one_chain_line_integral_def using 0 \finite one_chain\ by (force simp add: intro: comm_monoid_add_class.sum.mono_neutral_left) qed lemma boundary_chain_diff: assumes "boundary_chain one_chain" shows "boundary_chain (one_chain - s)" using assms by (auto simp add: boundary_chain_def) lemma gen_common_subdivision_imp_eq_line_integral: assumes "(common_sudiv_exists one_chain1 one_chain2)" "boundary_chain one_chain1" "boundary_chain one_chain2" "\(k, \)\one_chain1. line_integral_exists F basis \" "finite one_chain1" "finite one_chain2" "finite basis" shows "one_chain_line_integral F basis one_chain1 = one_chain_line_integral F basis one_chain2" "\(k, \)\one_chain2. line_integral_exists F basis \" proof - obtain ps1 ps2 where gen_subdiv: "(common_boundary_sudivision_exists (one_chain1 - ps1) (one_chain2 - ps2))" "(\(k, \)\ps1. point_path \)" " (\(k, \)\ps2. point_path \)" using assms(1) gen_common_subdiv_imp_common_subdiv by blast show "one_chain_line_integral F basis one_chain1 = one_chain_line_integral F basis one_chain2" using one_chain_line_integral_point_paths gen_common_subdiv_imp_common_subdiv assms(2-7) gen_subdiv common_subdivision_imp_eq_line_integral(1)[OF gen_subdiv(1) boundary_chain_diff[OF assms(2)] boundary_chain_diff[OF assms(3)]] by auto show "\(k, \)\one_chain2. line_integral_exists F basis \" proof- obtain subdiv where subdiv_props: "chain_subdiv_chain (one_chain1-ps1) subdiv" "chain_subdiv_chain (one_chain2-ps2) subdiv" "(\(k, \) \ subdiv. valid_path \)" "(boundary_chain subdiv)" using gen_subdiv(1) by (auto simp add: common_boundary_sudivision_exists_def) have "\(k, \)\subdiv. line_integral_exists F basis \" using one_chain_line_integral_eq_line_integral_on_sudivision'(2)[OF subdiv_props(1) boundary_chain_diff[OF assms(2)] subdiv_props(4)] assms(4) subdiv_props(3) assms(5) assms(7) by blast then have i: "\(k, \)\one_chain2-ps2. line_integral_exists F basis \" using one_chain_line_integral_eq_line_integral_on_sudivision(2)[OF subdiv_props(2) boundary_chain_diff[OF assms(3)] subdiv_props(4)] subdiv_props(3) assms(6) assms(7) by blast then show ?thesis using gen_subdiv(3) line_integral_exists_point_path[OF assms(7)] by blast qed qed lemma common_sudiv_exists_refl: assumes "common_sudiv_exists C1 C2" shows "common_sudiv_exists C2 C1" using assms apply(simp add: common_sudiv_exists_def) by auto lemma chain_subdiv_path_singleton: shows "chain_subdiv_path \ {(1,\)}" proof - have "rec_join [(1,\)] = \" by (simp add: joinpaths_def) then have "set [(1,\)] = {(1, \)}" "distinct [(1,\)]" "rec_join [(1,\)] = \" "valid_chain_list [(1,\)]" by auto then show ?thesis by (metis (no_types) chain_subdiv_path.intros) qed lemma chain_subdiv_path_singleton_reverse: shows "chain_subdiv_path (reversepath \) {(-1,\)}" proof - have "rec_join [(-1,\)] = reversepath \" by (simp add: joinpaths_def) then have "set [(-1,\)] = {(- 1, \)}" "distinct [(-1,\)]" "rec_join [(-1,\)] = reversepath \" "valid_chain_list [(-1,\)]" by auto then have "chain_subdiv_path (reversepath \) (set [(- 1, \)])" using chain_subdiv_path.intros by blast then show ?thesis by simp qed lemma chain_subdiv_chain_refl: assumes "boundary_chain C" shows "chain_subdiv_chain C C" using chain_subdiv_path_singleton chain_subdiv_path_singleton_reverse assms unfolding chain_subdiv_chain_def boundary_chain_def pairwise_def using case_prodI2 coeff_cube_to_path.simps by (rule_tac x="\x. {x}" in exI) auto (*path reparam_weaketrization*) definition reparam_weak where "reparam_weak \1 \2 \ \\. (\x\{0..1}. \1 x = (\2 o \) x) \ \ piecewise_C1_differentiable_on {0..1} \ \(0) = 0 \ \(1) = 1 \ \ ` {0..1} = {0..1}" definition reparam where "reparam \1 \2 \ \\. (\x\{0..1}. \1 x = (\2 o \) x) \ \ piecewise_C1_differentiable_on {0..1} \ \(0) = 0 \ \(1) = 1 \ bij_betw \ {0..1} {0..1} \ \ -` {0..1} \ {0..1} \ (\x\{0..1}. finite (\ -` {x}))" lemma reparam_weak_eq_refl: shows "reparam_weak \1 \1" unfolding reparam_weak_def apply (rule_tac x="id" in exI) by (auto simp add: id_def piecewise_C1_differentiable_on_def C1_differentiable_on_def continuous_on_id) lemma line_integral_exists_smooth_one_base: assumes "\ C1_differentiable_on {0..1}" (*To generalise this to valid_path we need a version of has_integral_substitution_strong that allows finite discontinuities of f*) "continuous_on (path_image \) (\x. F x \ b)" shows "line_integral_exists F {b} \" proof- have gamma2_differentiable: "(\x \ {0 .. 1}. \ differentiable at x)" using assms(1) by (auto simp add: valid_path_def C1_differentiable_on_eq) then have gamma2_b_component_differentiable: "(\x \ {0 .. 1}. (\x. (\ x) \ b) differentiable at x)" by auto then have "(\x. (\ x) \ b) differentiable_on {0..1}" using differentiable_at_withinI by (auto simp add: differentiable_on_def) then have gama2_cont_comp: "continuous_on {0..1} (\x. (\ x) \ b)" using differentiable_imp_continuous_on by auto have gamma2_cont:"continuous_on {0..1} \" using assms(1) C1_differentiable_imp_continuous_on by (auto simp add: valid_path_def) have iii: "continuous_on {0..1} (\x. F (\ x) \ b * (vector_derivative \ (at x within {0..1}) \ b))" proof- have 0: "continuous_on {0..1} (\x. F (\ x) \ b)" using assms(2) continuous_on_compose[OF gamma2_cont] by (auto simp add: path_image_def) obtain D where D: "(\x\{0..1}. (\ has_vector_derivative D x) (at x)) \ continuous_on {0..1} D" using assms(1) by (auto simp add: C1_differentiable_on_def) then have *:"\x\{0..1}. vector_derivative \ (at x within{0..1}) = D x" using vector_derivative_at vector_derivative_at_within_ivl by fastforce then have "continuous_on {0..1} (\x. vector_derivative \ (at x within{0..1}))" using continuous_on_eq D by force then have 1: "continuous_on {0..1} (\x. (vector_derivative \ (at x within{0..1})) \ b)" by(auto intro!: continuous_intros) show ?thesis using continuous_on_mult[OF 0 1] by auto qed then have "(\x. F (\ x) \ b * (vector_derivative \ (at x within {0..1}) \ b)) integrable_on {0..1}" using integrable_continuous_real by auto then show "line_integral_exists F {b} \" by(auto simp add: line_integral_exists_def) qed lemma contour_integral_primitive_lemma: fixes f :: "complex \ complex" and g :: "real \ complex" assumes "a \ b" and "\x. x \ s \ (f has_field_derivative f' x) (at x within s)" and "g piecewise_differentiable_on {a..b}" "\x. x \ {a..b} \ g x \ s" shows "((\x. f'(g x) * vector_derivative g (at x within {a..b})) has_integral (f(g b) - f(g a))) {a..b}" proof - obtain k where k: "finite k" "\x\{a..b} - k. g differentiable (at x within {a..b})" and cg: "continuous_on {a..b} g" using assms by (auto simp: piecewise_differentiable_on_def) have cfg: "continuous_on {a..b} (\x. f (g x))" apply (rule continuous_on_compose [OF cg, unfolded o_def]) using assms apply (metis field_differentiable_def field_differentiable_imp_continuous_at continuous_on_eq_continuous_within continuous_on_subset image_subset_iff) done { fix x::real assume a: "a < x" and b: "x < b" and xk: "x \ k" then have "g differentiable at x within {a..b}" using k by (simp add: differentiable_at_withinI) then have "(g has_vector_derivative vector_derivative g (at x within {a..b})) (at x within {a..b})" by (simp add: vector_derivative_works has_field_derivative_def scaleR_conv_of_real) then have gdiff: "(g has_derivative (\u. u * vector_derivative g (at x within {a..b}))) (at x within {a..b})" by (simp add: has_vector_derivative_def scaleR_conv_of_real) have "(f has_field_derivative (f' (g x))) (at (g x) within g ` {a..b})" using assms by (metis a atLeastAtMost_iff b DERIV_subset image_subset_iff less_eq_real_def) then have fdiff: "(f has_derivative (*) (f' (g x))) (at (g x) within g ` {a..b})" by (simp add: has_field_derivative_def) have "((\x. f (g x)) has_vector_derivative f' (g x) * vector_derivative g (at x within {a..b})) (at x within {a..b})" using diff_chain_within [OF gdiff fdiff] by (simp add: has_vector_derivative_def scaleR_conv_of_real o_def mult_ac) } note * = this show ?thesis apply (rule fundamental_theorem_of_calculus_interior_strong) using k assms cfg * apply (auto simp: at_within_Icc_at) done qed lemma line_integral_primitive_lemma: (*Only for real normed fields, i.e. vectors with products*) fixes f :: "'a::{euclidean_space,real_normed_field} \ 'a::{euclidean_space,real_normed_field}" and g :: "real \ 'a" assumes "\(a::'a). a \ s \ (f has_field_derivative (f' a)) (at a within s)" and "g piecewise_differentiable_on {0::real..1}" "\x. x \ {0..1} \ g x \ s" and "base_vec \ Basis" shows "((\x. ((f'(g x)) * (vector_derivative g (at x within {0..1}))) \ base_vec) has_integral (((f(g 1)) \ base_vec - (f(g 0)) \ base_vec))) {0..1}" proof - obtain k where k: "finite k" "\x\{0..1} - k. g differentiable (at x within {0..1})" and cg: "continuous_on {0..1} g" using assms by (auto simp: piecewise_differentiable_on_def) have cfg: "continuous_on {0..1} (\x. f (g x))" apply (rule continuous_on_compose [OF cg, unfolded o_def]) using assms apply (metis field_differentiable_def field_differentiable_imp_continuous_at continuous_on_eq_continuous_within continuous_on_subset image_subset_iff) done { fix x::real assume a: "0 < x" and b: "x < 1" and xk: "x \ k" then have "g differentiable at x within {0..1}" using k by (simp add: differentiable_at_withinI) then have "(g has_vector_derivative vector_derivative g (at x within {0..1})) (at x within {0..1})" by (simp add: vector_derivative_works has_field_derivative_def scaleR_conv_of_real) then have gdiff: "(g has_derivative (\u. of_real u * vector_derivative g (at x within {0..1}))) (at x within {0..1})" by (simp add: has_vector_derivative_def scaleR_conv_of_real) have "(f has_field_derivative (f' (g x))) (at (g x) within g ` {0..1})" using assms by (metis a atLeastAtMost_iff b DERIV_subset image_subset_iff less_eq_real_def) then have fdiff: "(f has_derivative (*) (f' (g x))) (at (g x) within g ` {0..1})" by (simp add: has_field_derivative_def) have "((\x. f (g x)) has_vector_derivative f' (g x) * vector_derivative g (at x within {0..1})) (at x within {0..1})" using diff_chain_within [OF gdiff fdiff] by (simp add: has_vector_derivative_def scaleR_conv_of_real o_def mult_ac) } then have *: "\x. x\{0<..<1} - k \ ((\x. f (g x)) has_vector_derivative f' (g x) * vector_derivative g (at x within {0..1})) (at x within {0..1})" by auto have "((\x. ((f'(g x))) * ((vector_derivative g (at x within {0..1})))) has_integral (((f(g 1)) - (f(g 0))))) {0..1}" using fundamental_theorem_of_calculus_interior_strong[OF k(1) zero_le_one _ cfg] using k assms cfg * by (auto simp: at_within_Icc_at) then have "((\x. (((f'(g x))) * ((vector_derivative g (at x within {0..1})))) \ base_vec) has_integral (((f(g 1)) - (f(g 0)))) \ base_vec) {0..1}" using has_integral_componentwise_iff assms(4) by fastforce then show ?thesis using inner_mult_left' by (simp add: inner_mult_left' inner_diff_left) qed lemma reparam_eq_line_integrals: assumes reparam: "reparam \1 \2" and pw_smooth: "\2 piecewise_C1_differentiable_on {0..1}" and (*To generalise this to valid_path we need a version of has_integral_substitution_strong that allows finite discontinuities of f*) cont: "continuous_on (path_image \2) (\x. F x \ b)" and line_integral_ex: "line_integral_exists F {b} \2" (*We need to remove this and work on special cases like conservative fields and field/line combinations that satisfy the improper integrals theorem assumptions*) shows "line_integral F {b} \1 = line_integral F {b} \2" "line_integral_exists F {b} \1" proof- obtain \ where phi: "(\x\{0..1}. \1 x = (\2 o \) x)"" \ piecewise_C1_differentiable_on {0..1}"" \(0) = 0"" \(1) = 1""bij_betw \ {0..1} {0..1}" "\ -` {0..1} \ {0..1}" "\x\{0..1}. finite (\ -` {x})" using reparam by (auto simp add: reparam_def) obtain s where s: "finite s" "\ C1_differentiable_on {0..1} - s" using phi by(auto simp add: reparam_def piecewise_C1_differentiable_on_def) let ?s = "s \ {0..1}" have s_inter: "finite ?s" "\ C1_differentiable_on {0..1} - ?s" using s apply blast by (metis Diff_Compl Diff_Diff_Int Diff_eq inf.commute s(2)) have cont_phi: "continuous_on {0..1} \" using phi by(auto simp add: reparam_def piecewise_C1_differentiable_on_imp_continuous_on) obtain s' D where s'_D: "finite s'" "(\x \ {0 .. 1} - s'. \2 differentiable at x)" "(\x\{0..1} - s'. (\2 has_vector_derivative D x) (at x)) \ continuous_on ({0..1} - s') D" using pw_smooth apply (auto simp add: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) by (simp add: vector_derivative_works) let ?s' = "s' \ {0..1}" have gamma2_differentiable: "finite ?s'" "(\x \ {0 .. 1} - ?s'. \2 differentiable at x)" "(\x\{0..1} - ?s'. (\2 has_vector_derivative D x) (at x)) \ continuous_on ({0..1} - ?s') D" using s'_D apply blast using s'_D(2) apply auto[1] by (metis Diff_Int2 inf_top.left_neutral s'_D(3)) then have gamma2_b_component_differentiable: "(\x \ {0 .. 1} - ?s'. (\x. (\2 x) \ b) differentiable at x)" using differentiable_inner by force then have "(\x. (\2 x) \ b) differentiable_on {0..1} - ?s'" using differentiable_at_withinI by (auto simp add: differentiable_on_def) then have gama2_cont_comp: "continuous_on ({0..1} - ?s') (\x. (\2 x) \ b)" using differentiable_imp_continuous_on by auto (**********Additional needed assumptions ************) have s_in01: "?s \ {0..1}" by auto have s'_in01: "?s' \ {0..1}" by auto have phi_backimg_s': "\ -` {0..1} \ {0..1}" using phi by auto have "inj_on \ {0..1}" using phi(5) by (auto simp add: bij_betw_def) have bij_phi: "bij_betw \ {0..1} {0..1}" using phi(5) by auto have finite_bck_img_single: "\x\{0..1}. finite (\ -` {x})" using phi by auto then have finite_bck_img_single_s': " \x\?s'. finite (\ -` {x})" by auto have gamma2_line_integrable: "(\x. F (\2 x) \ b * (vector_derivative \2 (at x within {0..1}) \ b)) integrable_on {0..1}" using line_integral_ex by (simp add: line_integral_exists_def) (****************************************************************) have finite_neg_img: "finite (\ -` ?s')" using finite_bck_img_single by (metis Int_iff finite_Int gamma2_differentiable(1) image_vimage_eq inf_img_fin_dom') have gamma2_cont:"continuous_on ({0..1} - ?s') \2" using gamma2_differentiable by (meson continuous_at_imp_continuous_on differentiable_imp_continuous_within) have iii: "continuous_on ({0..1} - ?s') (\x. F (\2 x) \ b * (vector_derivative \2 (at x within {0..1}) \ b))" proof- have 0: "continuous_on ({0..1} - ?s') (\x. F (\2 x) \ b)" using cont continuous_on_compose[OF gamma2_cont] continuous_on_compose2 gamma2_cont unfolding path_image_def by fastforce have D: "(\x\{0..1} - ?s'. (\2 has_vector_derivative D x) (at x)) \ continuous_on ({0..1} - ?s') D" using gamma2_differentiable by auto then have *:"\x\{0..1} - ?s'. vector_derivative \2 (at x within{0..1}) = D x" using vector_derivative_at vector_derivative_at_within_ivl by fastforce then have "continuous_on ({0..1} - ?s') (\x. vector_derivative \2 (at x within{0..1}))" using continuous_on_eq D by metis then have 1: "continuous_on ({0..1} - ?s') (\x. (vector_derivative \2 (at x within{0..1})) \ b)" by (auto intro!: continuous_intros) show ?thesis using continuous_on_mult[OF 0 1] by auto qed have iv: "\(0) \ \(1)" using phi(3) phi(4) by (simp add: reparam_def) have v: "\`{0..1} \ {0..1}" using phi by (auto simp add: reparam_def bij_betw_def) obtain D where D_props: "(\x\{0..1} - ?s. (\ has_vector_derivative D x) (at x))" using s by (auto simp add: C1_differentiable_on_def) then have "(\x. x \ ({0..1} - ?s) \ (\ has_vector_derivative D x) (at x within {0..1}))" using has_vector_derivative_at_within by blast then have vi: "(\x. x \ ({0..1} - ?s) \ (\ has_real_derivative D x) (at x within {0..1}))" using has_field_derivative_iff_has_vector_derivative by blast have a:"((\x. D x * (F (\2 (\ x)) \ b * (vector_derivative \2 (at (\ x) within {0..1}) \ b))) has_integral integral {\ 0..\ 1} (\x. F (\2 x) \ b * (vector_derivative \2 (at x within {0..1}) \ b))) ({0..1})" proof- have a: "integral {\ 1..\ 0} (\x. F (\2 x) \ b * (vector_derivative \2 (at x within {0..1}) \ b)) = 0" using integral_singleton integral_empty iv by (simp add: phi(3) phi(4)) have b: "((\x. D x *\<^sub>R (F (\2 (\ x)) \ b * (vector_derivative \2 (at (\ x) within {0..1}) \ b))) has_integral integral {\ 0..\ 1} (\x. F (\2 x) \ b * (vector_derivative \2 (at x within {0..1}) \ b)) - integral {\ 1..\ 0} (\x. F (\2 x) \ b * (vector_derivative \2 (at x within {0..1}) \ b))) {0..1}" apply(rule has_integral_substitution_general_'[OF s_inter(1) zero_le_one gamma2_differentiable(1) v gamma2_line_integrable iii cont_phi finite_bck_img_single_s']) proof- have "surj_on {0..1} \" using bij_phi by (metis (full_types) bij_betw_def image_subsetI rangeI) then show "surj_on ?s' \" using bij_phi s'_in01 by blast show "inj_on \ ({0..1} \ (?s \ \ -` ?s'))" proof- have i: "inj_on \ {0..1}" using bij_phi using bij_betw_def by blast have ii: "({0..1} \ (?s \ \ -` ?s')) = {0..1}" using phi_backimg_s' s_in01 s'_in01 by blast show ?thesis using i ii by auto qed show "\x. x \ {0..1} - ?s \ (\ has_real_derivative D x) (at x within {0..1})" using vi by blast qed show ?thesis using a b by auto qed then have b: "integral {0..1} (\x. D x * (F (\2 (\ x)) \ b * (vector_derivative \2 (at (\ x) within {0..1}) \ b))) = integral {\ 0..\ 1} (\x. F (\2 x) \ b * (vector_derivative \2 (at x within {0..1}) \ b))" by auto have gamma2_vec_diffable: "\x::real. x \ {0..1} - ((\ -` ?s') \ ?s) \ ((\2 o \) has_vector_derivative vector_derivative (\2 \ \) (at x)) (at x)" proof- fix x::real assume ass: "x \ {0..1} - ((\ -` ?s') \ ?s)" have zer_le_x_le_1:"0\ x \ x \ 1" using ass by simp show "((\2 o \) has_vector_derivative vector_derivative (\2 \ \) (at x)) (at x)" proof- have **: "\2 differentiable at (\ x)" using gamma2_differentiable(2) ass v by blast have ***: " \ differentiable at x" using s ass by (auto simp add: C1_differentiable_on_eq) then show "((\2 o \) has_vector_derivative vector_derivative (\2 \ \) (at x)) (at x)" using differentiable_chain_at[OF *** **] by (auto simp add: vector_derivative_works) qed qed then have gamma2_vec_deriv_within: "\x::real. x \ {0..1} - ((\ -` ?s') \ ?s) \ vector_derivative (\2 \ \) (at x) = vector_derivative (\2 \ \) (at x within {0..1})" using vector_derivative_at_within_ivl[OF gamma2_vec_diffable] by auto have "\x\{0..1} - ((\ -` ?s') \ ?s). D x * (vector_derivative \2 (at (\ x) within {0..1}) \ b) = (vector_derivative (\2 \ \) (at x within {0..1}) \ b)" proof fix x::real assume ass: "x \ {0..1} -((\ -` ?s') \ ?s)" then have 0: "\ differentiable (at x)" using s by (auto simp add: C1_differentiable_on_def differentiable_def has_vector_derivative_def) obtain D2 where "(\ has_vector_derivative D2) (at x)" using D_props ass by blast have "\ x \ {0..1} - ?s'" using phi(5) ass by (metis Diff_Un Diff_iff Int_iff bij_betw_def image_eqI vimageI) then have 1: "\2 differentiable (at (\ x))" using gamma2_differentiable by auto have 3:" vector_derivative \2 (at (\ x)) = vector_derivative \2 (at (\ x) within {0..1})" proof- have *:"0\ \ x \ \ x \ 1" using phi(5) ass using \\ x \ {0..1} - s' \ {0..1}\ by auto then have **:"(\2 has_vector_derivative (vector_derivative \2 (at (\ x)))) (at (\ x))" using 1 vector_derivative_works by auto show ?thesis using * vector_derivative_at_within_ivl[OF **] by auto qed show "D x * (vector_derivative \2 (at (\ x) within {0..1}) \ b) = vector_derivative (\2 \ \) (at x within {0..1}) \ b" using vector_derivative_chain_at[OF 0 1] apply (auto simp add: gamma2_vec_deriv_within[OF ass, symmetric] 3[symmetric]) using D_props ass vector_derivative_at by fastforce qed then have c:"\x. x\({0..1} -((\ -` ?s') \ ?s)) \ D x * (F (\2 (\ x)) \ b * (vector_derivative \2 (at (\ x) within {0..1}) \ b)) = F (\2 (\ x)) \ b * (vector_derivative (\2 \ \) (at x within {0..1}) \ b)" by auto then have d: "integral ({0..1}) (\x. D x * (F (\2 (\ x)) \ b * (vector_derivative \2 (at (\ x) within {0..1}) \ b))) = integral ({0..1}) (\x. F (\2 (\ x)) \ b * (vector_derivative (\2 \ \) (at x within {0..1}) \ b))" proof- have "negligible ((\ -` ?s') \ ?s)" using finite_neg_img s(1) by auto then show ?thesis using c integral_spike by (metis(no_types,lifting)) qed have phi_in_int: "(\x. x \ {0..1} \ \ x \ {0..1})" using phi using v by blast then have e: "((\x. F (\2 (\ x)) \ b * (vector_derivative (\2 \ \) (at x within {0..1}) \ b)) has_integral integral {\ 0..\ 1} (\x. F (\2 x) \ b * (vector_derivative \2 (at x within {0..1}) \ b))){0..1}" proof- have "negligible ?s" using s_inter(1) by auto have 0: "negligible ((\ -` ?s') \ ?s)" using finite_neg_img s(1) by auto have c':"\ x\ {0..1} - ((\ -` ?s') \ ?s). D x * (F (\2 (\ x)) \ b * (vector_derivative \2 (at (\ x) within {0..1}) \ b)) = F (\2 (\ x)) \ b * (vector_derivative (\2 \ \) (at x within {0..1}) \ b)" using c by blast have has_integral_spike_eq': "\s t f g y. negligible s \ \x\t - s. g x = f x \ (f has_integral y) t = (g has_integral y) t" using has_integral_spike_eq by metis show ?thesis using a has_integral_spike_eq'[OF 0 c'] by blast qed then have f: "((\x. F (\1 x) \ b * (vector_derivative \1 (at x within {0..1}) \ b)) has_integral integral {\ 0..\ 1} (\x. F (\2 x) \ b * (vector_derivative \2 (at x within {0..1}) \ b))) {0..1}" proof- assume ass: "((\x. F (\2 (\ x)) \ b * (vector_derivative (\2 \ \) (at x within {0..1}) \ b)) has_integral integral {\ 0..\ 1} (\x. F (\2 x) \ b * (vector_derivative \2 (at x within {0..1}) \ b))) {0..1}" have *:"\x\{0..1} - ( ((\ -` ?s') \ ?s) \ {0,1}). (\x. F (\2 (\ x)) \ b * (vector_derivative (\2 \ \) (at x within {0..1}) \ b)) x = (\x. F (\1 x) \ b * (vector_derivative \1 (at x within {0..1}) \ b)) x" proof- have "\x\{0<..<1} - (\ -` ?s' \ ?s). (vector_derivative (\2 \ \) (at x within {0..1}) \ b) = (vector_derivative (\1) (at x within {0..1}) \ b)" proof have i: "open ({0<..<1} - ((\ -` ?s') \ ?s))" using open_diff s(1) open_greaterThanLessThan finite_neg_img by (simp add: open_diff) have ii: "\x\{0<..<1::real} - (\ -` ?s' \ ?s). (\2 \ \) x = \1 x" using phi(1) by auto fix x::real assume ass: " x \ {0<..<1::real} - ((\ -` ?s') \ ?s)" then have iii: "(\2 \ \ has_vector_derivative vector_derivative (\2 \ \) (at x within {0..1})) (at x)" by (metis (no_types) Diff_iff add.commute add_strict_mono ass atLeastAtMost_iff gamma2_vec_deriv_within gamma2_vec_diffable greaterThanLessThan_iff less_irrefl not_le) (*Most crucial but annoying step*) then have iv:"(\1 has_vector_derivative vector_derivative (\2 \ \) (at x within {0..1})) (at x)" using has_derivative_transform_within_open i ii ass apply(auto simp add: has_vector_derivative_def) apply (meson ass has_derivative_transform_within_open i ii) apply (meson ass has_derivative_transform_within_open i ii) by (meson ass has_derivative_transform_within_open i ii) have v: "0 \ x" "x \ 1" using ass by auto have 0: "vector_derivative \1 (at x within {0..1}) = vector_derivative (\2 \ \) (at x within {0..1})" using vector_derivative_at_within_ivl[OF iv v(1) v(2) zero_less_one] by force have 1: "vector_derivative (\2 \ \) (at x within {0..1}) = vector_derivative (\2 \ \) (at x within {0..1})" using vector_derivative_at_within_ivl[OF iii v(1) v(2) zero_less_one] by force then have "vector_derivative (\2 \ \) (at x within {0..1}) = vector_derivative \1 (at x within {0..1})" using 0 1 by auto then show "vector_derivative (\2 \ \) (at x within {0..1}) \ b = vector_derivative \1 (at x within {0..1}) \ b" by auto qed then have i: "\x\{0..1} - ( ((\ -` ?s') \ ?s)\{0,1}). (vector_derivative (\2 \ \) (at x within {0..1}) \ b) = (vector_derivative (\1) (at x within {0..1}) \ b)" by auto have ii: "\x\{0..1} - (((\ -` ?s') \ ?s)\{0,1}). F (\1 x) \ b = F (\2 (\ x)) \ b" using phi(1) by auto show ?thesis using i ii by metis qed have **: "negligible ((\ -` ?s') \ ?s \ {0, 1})" using s(1) finite_neg_img by auto have has_integral_spike_eq': "\s t g f y. negligible s \ \x\t - s. g x = f x \ (f has_integral y) t = (g has_integral y) t" using has_integral_spike_eq by metis show ?thesis using has_integral_spike_eq'[OF ** *] ass by blast qed then show "line_integral_exists F {b} \1" using phi by(auto simp add: line_integral_exists_def) have "integral ({0..1}) (\x. F (\2 (\ x)) \ b * (vector_derivative (\2 \ \) (at x within {0..1}) \ b)) = integral ({0..1}) (\x. F (\1 x) \ b * (vector_derivative \1 (at x within {0..1}) \ b))" using integral_unique[OF e] integral_unique[OF f] by metis moreover have "integral ({0..1}) (\x. F (\2 (\ x)) \ b * (vector_derivative (\2 \ \) (at x within {0..1}) \ b)) = integral ({0..1}) (\x. F (\2 x) \ b * (vector_derivative \2 (at x within {0..1}) \ b))" using b d phi by (auto simp add:) ultimately show "line_integral F {b} \1 = line_integral F {b} \2" using phi by(auto simp add: line_integral_def) qed lemma reparam_weak_eq_line_integrals: assumes "reparam_weak \1 \2" "\2 C1_differentiable_on {0..1}" (*To generalise this to valid_path we need a version of has_integral_substitution_strong that allows finite discontinuities of f*) "continuous_on (path_image \2) (\x. F x \ b)" shows "line_integral F {b} \1 = line_integral F {b} \2" "line_integral_exists F {b} \1" proof- obtain \ where phi: "(\x\{0..1}. \1 x = (\2 o \) x)"" \ piecewise_C1_differentiable_on {0..1}"" \(0) = 0"" \(1) = 1"" \ ` {0..1} = {0..1}" using assms(1) by (auto simp add: reparam_weak_def) obtain s where s: "finite s" "\ C1_differentiable_on {0..1} - s" using phi by(auto simp add: reparam_weak_def piecewise_C1_differentiable_on_def) have cont_phi: "continuous_on {0..1} \" using phi by(auto simp add: reparam_weak_def piecewise_C1_differentiable_on_imp_continuous_on) have gamma2_differentiable: "(\x \ {0 .. 1}. \2 differentiable at x)" using assms(2) by (auto simp add: valid_path_def C1_differentiable_on_eq) then have gamma2_b_component_differentiable: "(\x \ {0 .. 1}. (\x. (\2 x) \ b) differentiable at x)" by auto then have "(\x. (\2 x) \ b) differentiable_on {0..1}" using differentiable_at_withinI by (auto simp add: differentiable_on_def) then have gama2_cont_comp: "continuous_on {0..1} (\x. (\2 x) \ b)" using differentiable_imp_continuous_on by auto have gamma2_cont:"continuous_on {0..1} \2" using assms(2) C1_differentiable_imp_continuous_on by (auto simp add: valid_path_def) have iii: "continuous_on {0..1} (\x. F (\2 x) \ b * (vector_derivative \2 (at x within {0..1}) \ b))" proof- have 0: "continuous_on {0..1} (\x. F (\2 x) \ b)" using assms(3) continuous_on_compose[OF gamma2_cont] by (auto simp add: path_image_def) obtain D where D: "(\x\{0..1}. (\2 has_vector_derivative D x) (at x)) \ continuous_on {0..1} D" using assms(2) by (auto simp add: C1_differentiable_on_def) then have *:"\x\{0..1}. vector_derivative \2 (at x within{0..1}) = D x" using vector_derivative_at vector_derivative_at_within_ivl by fastforce then have "continuous_on {0..1} (\x. vector_derivative \2 (at x within{0..1}))" using continuous_on_eq D by force then have 1: "continuous_on {0..1} (\x. (vector_derivative \2 (at x within{0..1})) \ b)" by (auto intro!: continuous_intros) show ?thesis using continuous_on_mult[OF 0 1] by auto qed have iv: "\(0) \ \(1)" using phi(3) phi(4) by (simp add: reparam_weak_def) have v: "\`{0..1} \ {0..1}" using phi(5) by (simp add: reparam_weak_def) obtain D where D_props: "(\x\{0..1} - s. (\ has_vector_derivative D x) (at x))" using s by (auto simp add: C1_differentiable_on_def) then have "(\x. x \ ({0..1} -s) \ (\ has_vector_derivative D x) (at x within {0..1}))" using has_vector_derivative_at_within by blast then have vi: "(\x. x \ ({0..1} - s) \ (\ has_real_derivative D x) (at x within {0..1}))" using has_field_derivative_iff_has_vector_derivative by blast have a:"((\x. D x * (F (\2 (\ x)) \ b * (vector_derivative \2 (at (\ x) within {0..1}) \ b))) has_integral integral {\ 0..\ 1} (\x. F (\2 x) \ b * (vector_derivative \2 (at x within {0..1}) \ b))) {0..1}" using has_integral_substitution_strong[OF s(1) zero_le_one iv v iii cont_phi vi] by simp then have b: "integral {0..1} (\x. D x * (F (\2 (\ x)) \ b * (vector_derivative \2 (at (\ x) within {0..1}) \ b))) = integral {\ 0..\ 1} (\x. F (\2 x) \ b * (vector_derivative \2 (at x within {0..1}) \ b))" by auto have gamma2_vec_diffable: "\x::real. x \ {0..1} -s \ ((\2 o \) has_vector_derivative vector_derivative (\2 \ \) (at x)) (at x)" proof- fix x::real assume ass: "x \ {0..1} -s" have zer_le_x_le_1:"0\ x \ x \ 1" using ass by auto show "((\2 o \) has_vector_derivative vector_derivative (\2 \ \) (at x)) (at x)" proof- have **: "\2 differentiable at (\ x)" using phi gamma2_differentiable by (auto simp add: zer_le_x_le_1) have ***: " \ differentiable at x" using s ass by (auto simp add: C1_differentiable_on_eq) then show "((\2 o \) has_vector_derivative vector_derivative (\2 \ \) (at x)) (at x)" using differentiable_chain_at[OF *** **] by (auto simp add: vector_derivative_works) qed qed then have gamma2_vec_deriv_within: "\x::real. x \ {0..1} -s \ vector_derivative (\2 \ \) (at x) = vector_derivative (\2 \ \) (at x within {0..1})" using vector_derivative_at_within_ivl[OF gamma2_vec_diffable] by auto have "\x\{0..1} - s. D x * (vector_derivative \2 (at (\ x) within {0..1}) \ b) = (vector_derivative (\2 \ \) (at x within {0..1}) \ b)" proof fix x::real assume ass: "x \ {0..1} -s" then have 0: "\ differentiable (at x)" using s by (auto simp add: C1_differentiable_on_def differentiable_def has_vector_derivative_def) obtain D2 where "(\ has_vector_derivative D2) (at x)" using D_props ass by blast have "\ x \ {0..1}" using phi(5) ass by (auto simp add: reparam_weak_def) then have 1: "\2 differentiable (at (\ x))" using gamma2_differentiable by auto have 3:" vector_derivative \2 (at (\ x)) = vector_derivative \2 (at (\ x) within {0..1})" proof- have *:"0\ \ x \ \ x \ 1" using phi(5) ass by auto then have **:"(\2 has_vector_derivative (vector_derivative \2 (at (\ x)))) (at (\ x))" using 1 vector_derivative_works by auto show ?thesis using * vector_derivative_at_within_ivl[OF **] by auto qed show "D x * (vector_derivative \2 (at (\ x) within {0..1}) \ b) = vector_derivative (\2 \ \) (at x within {0..1}) \ b" using vector_derivative_chain_at[OF 0 1] apply (auto simp add: gamma2_vec_deriv_within[OF ass, symmetric] 3[symmetric]) using D_props ass vector_derivative_at by fastforce qed then have c:"\x. x\({0..1} -s) \ D x * (F (\2 (\ x)) \ b * (vector_derivative \2 (at (\ x) within {0..1}) \ b)) = F (\2 (\ x)) \ b * (vector_derivative (\2 \ \) (at x within {0..1}) \ b)" by auto then have d: "integral ({0..1}) (\x. D x * (F (\2 (\ x)) \ b * (vector_derivative \2 (at (\ x) within {0..1}) \ b))) = integral ({0..1}) (\x. F (\2 (\ x)) \ b * (vector_derivative (\2 \ \) (at x within {0..1}) \ b))" proof- have "negligible s" using s(1) by auto then show ?thesis using c integral_spike by (metis(no_types,lifting)) qed have phi_in_int: "(\x. x \ {0..1} \ \ x \ {0..1})" using phi by(auto simp add:) then have e: "((\x. F (\2 (\ x)) \ b * (vector_derivative (\2 \ \) (at x within {0..1}) \ b)) has_integral integral {\ 0..\ 1} (\x. F (\2 x) \ b * (vector_derivative \2 (at x within {0..1}) \ b))){0..1}" proof- have 0:"negligible s" using s(1) by auto have c':"\ x\ {0..1} -s. D x * (F (\2 (\ x)) \ b * (vector_derivative \2 (at (\ x) within {0..1}) \ b)) = F (\2 (\ x)) \ b * (vector_derivative (\2 \ \) (at x within {0..1}) \ b)" using c by auto have has_integral_spike_eq': "\s t f g y. negligible s \ \x\t - s. g x = f x \ (f has_integral y) t = (g has_integral y) t" using has_integral_spike_eq by metis show ?thesis using a has_integral_spike_eq'[OF 0 c'] by blast qed then have f: "((\x. F (\1 x) \ b * (vector_derivative \1 (at x within {0..1}) \ b)) has_integral integral {\ 0..\ 1} (\x. F (\2 x) \ b * (vector_derivative \2 (at x within {0..1}) \ b))) {0..1}" proof- assume ass: "((\x. F (\2 (\ x)) \ b * (vector_derivative (\2 \ \) (at x within {0..1}) \ b)) has_integral integral {\ 0..\ 1} (\x. F (\2 x) \ b * (vector_derivative \2 (at x within {0..1}) \ b))) {0..1}" have *:"\x\{0..1} - (s\{0,1}). (\x. F (\2 (\ x)) \ b * (vector_derivative (\2 \ \) (at x within {0..1}) \ b)) x = (\x. F (\1 x) \ b * (vector_derivative \1 (at x within {0..1}) \ b)) x" proof- have "\x\{0<..<1} - s. (vector_derivative (\2 \ \) (at x within {0..1}) \ b) = (vector_derivative (\1) (at x within {0..1}) \ b)" proof have i: "open ({0<..<1} - s)" using open_diff s open_greaterThanLessThan by blast have ii: "\x\{0<..<1::real} - s. (\2 \ \) x = \1 x" using phi(1) by auto fix x::real assume ass: " x \ {0<..<1::real} - s" then have iii: "(\2 \ \ has_vector_derivative vector_derivative (\2 \ \) (at x within {0..1})) (at x)" using has_vector_derivative_at_within gamma2_vec_deriv_within gamma2_vec_diffable by auto (*Most crucial but annoying step*) then have iv:"(\1 has_vector_derivative vector_derivative (\2 \ \) (at x within {0..1})) (at x)" using has_derivative_transform_within_open i ii ass apply(auto simp add: has_vector_derivative_def) by force have v: "0 \ x" "x \ 1" using ass by auto have 0: "vector_derivative \1 (at x within {0..1}) = vector_derivative (\2 \ \) (at x within {0..1})" using vector_derivative_at_within_ivl[OF iv v(1) v(2) zero_less_one] by force have 1: "vector_derivative (\2 \ \) (at x within {0..1}) = vector_derivative (\2 \ \) (at x within {0..1})" using vector_derivative_at_within_ivl[OF iii v(1) v(2) zero_less_one] by force then have "vector_derivative (\2 \ \) (at x within {0..1}) = vector_derivative \1 (at x within {0..1})" using 0 1 by auto then show "vector_derivative (\2 \ \) (at x within {0..1}) \ b = vector_derivative \1 (at x within {0..1}) \ b" by auto qed then have i: "\x\{0..1} - (s\{0,1}). (vector_derivative (\2 \ \) (at x within {0..1}) \ b) = (vector_derivative (\1) (at x within {0..1}) \ b)" by auto have ii: "\x\{0..1} - (s\{0,1}). F (\1 x) \ b = F (\2 (\ x)) \ b" using phi(1) by auto show ?thesis using i ii by auto qed have **: "negligible (s\{0,1})" using s(1) by auto have has_integral_spike_eq': "\s t g f y. negligible s \ \x\t - s. g x = f x \ (f has_integral y) t = (g has_integral y) t" using has_integral_spike_eq by metis show ?thesis using has_integral_spike_eq'[OF ** *] ass by blast qed then show "line_integral_exists F {b} \1" using phi by(auto simp add: line_integral_exists_def) have "integral ({0..1}) (\x. F (\2 (\ x)) \ b * (vector_derivative (\2 \ \) (at x within {0..1}) \ b)) = integral ({0..1}) (\x. F (\1 x) \ b * (vector_derivative \1 (at x within {0..1}) \ b))" using integral_unique[OF e] integral_unique[OF f] by metis moreover have "integral ({0..1}) (\x. F (\2 (\ x)) \ b * (vector_derivative (\2 \ \) (at x within {0..1}) \ b)) = integral ({0..1}) (\x. F (\2 x) \ b * (vector_derivative \2 (at x within {0..1}) \ b))" using b d phi by (auto simp add:) ultimately show "line_integral F {b} \1 = line_integral F {b} \2" using phi by(auto simp add: line_integral_def) qed lemma line_integral_sum_basis: assumes "finite (basis::('a::euclidean_space) set)" "\b\basis. line_integral_exists F {b} \" shows "line_integral F basis \ = (\b\basis. line_integral F {b} \)" "line_integral_exists F basis \" using assms proof(induction basis) show "line_integral F {} \ = (\b\{}. line_integral F {b} \)" by(auto simp add: line_integral_def) show "\b\{}. line_integral_exists F {b} \ \ line_integral_exists F {} \" by(simp add: line_integral_exists_def integrable_0) next fix basis::"('a::euclidean_space) set" fix x::"'a::euclidean_space" fix \ assume ind_hyp: "(\b\basis. line_integral_exists F {b} \ \ line_integral_exists F basis \)" "(\b\basis. line_integral_exists F {b} \ \ line_integral F basis \ = (\b\basis. line_integral F {b} \))" assume step: "finite basis" "x \ basis" "\b\insert x basis. line_integral_exists F {b} \" then have 0: "line_integral_exists F {x} \" by auto have 1:"line_integral_exists F basis \" using ind_hyp step by auto then show "line_integral_exists F (insert x basis) \" using step(1) step(2) line_integral_sum_gen(2)[OF _ 0 1] by simp have 3: "finite (insert x basis)" using step(1) by auto have "line_integral F basis \ = (\b\basis. line_integral F {b} \)" using ind_hyp step by auto then show "line_integral F (insert x basis) \ = (\b\insert x basis. line_integral F {b} \)" using step(1) step(2) line_integral_sum_gen(1)[OF 3 0 1] by force qed lemma reparam_weak_eq_line_integrals_basis: assumes "reparam_weak \1 \2" "\2 C1_differentiable_on {0..1}" (*To generalise this to valid_path we need a version of has_integral_substitution_strong that allows finite discontinuities*) "\b\basis. continuous_on (path_image \2) (\x. F x \ b)" "finite basis" shows "line_integral F basis \1 = line_integral F basis \2" "line_integral_exists F basis \1" proof- show "line_integral_exists F basis \1" using reparam_weak_eq_line_integrals(2)[OF assms(1) assms(2)] assms(3-4) line_integral_sum_basis(2)[OF assms(4)] by(simp add: subset_iff) show "line_integral F basis \1 = line_integral F basis \2" using reparam_weak_eq_line_integrals[OF assms(1) assms(2)] assms(3-4) line_integral_sum_basis(1)[OF assms(4)] line_integral_exists_smooth_one_base[OF assms(2)] by(simp add: subset_iff) qed lemma reparam_eq_line_integrals_basis: assumes "reparam \1 \2" "\2 piecewise_C1_differentiable_on {0..1}" "\b\basis. continuous_on (path_image \2) (\x. F x \ b)" "finite basis" "\b\basis. line_integral_exists F {b} \2" (*We need to remove this and work on special cases like conservative fields and field/line combinations that satisfy the improper integrals theorem assumptions*) shows "line_integral F basis \1 = line_integral F basis \2" "line_integral_exists F basis \1" proof- show "line_integral_exists F basis \1" using reparam_eq_line_integrals(2)[OF assms(1) assms(2)] assms(3-5) line_integral_sum_basis(2)[OF assms(4)] by(simp add: subset_iff) show "line_integral F basis \1 = line_integral F basis \2" using reparam_eq_line_integrals[OF assms(1) assms(2)] assms(3-5) line_integral_sum_basis(1)[OF assms(4)] by(simp add: subset_iff) qed lemma line_integral_exists_smooth: assumes "\ C1_differentiable_on {0..1}" (*To generalise this to valid_path we need a version of has_integral_substitution_strong that allows finite discontinuities*) "\(b::'a::euclidean_space) \basis. continuous_on (path_image \) (\x. F x \ b)" "finite basis" shows "line_integral_exists F basis \" using reparam_weak_eq_line_integrals_basis(2)[OF reparam_weak_eq_refl[where ?\1.0 = \]] assms by fastforce lemma smooth_path_imp_reverse: assumes "g C1_differentiable_on {0..1}" shows "(reversepath g) C1_differentiable_on {0..1}" using assms continuous_on_const apply (auto simp: reversepath_def) apply (rule C1_differentiable_compose [of "\x::real. 1-x" _ g, unfolded o_def]) apply (auto simp: C1_differentiable_on_eq) apply (simp add: finite_vimageI inj_on_def) done lemma piecewise_smooth_path_imp_reverse: assumes "g piecewise_C1_differentiable_on {0..1}" shows "(reversepath g) piecewise_C1_differentiable_on {0..1}" using assms valid_path_reversepath using valid_path_def by blast definition chain_reparam_weak_chain where "chain_reparam_weak_chain one_chain1 one_chain2 \ \f. bij f \ f ` one_chain1 = one_chain2 \ (\(k,\)\one_chain1. if k = fst (f(k,\)) then reparam_weak \ (snd (f(k,\))) else reparam_weak \ (reversepath (snd (f(k,\)))))" lemma chain_reparam_weak_chain_line_integral: assumes "chain_reparam_weak_chain one_chain1 one_chain2" "\(k2,\2)\one_chain2. \2 C1_differentiable_on {0..1}" "\(k2,\2)\one_chain2.\b\basis. continuous_on (path_image \2) (\x. F x \ b)" "finite basis" and bound1: "boundary_chain one_chain1" and bound2: "boundary_chain one_chain2" shows "one_chain_line_integral F basis one_chain1 = one_chain_line_integral F basis one_chain2" "\(k, \)\one_chain1. line_integral_exists F basis \" proof- obtain f where f: "bij f" "(\(k,\)\one_chain1. if k = fst (f(k,\)) then reparam_weak \ (snd (f(k,\))) else reparam_weak \ (reversepath (snd (f(k,\)))))" "f ` one_chain1 = one_chain2" using assms(1) by (auto simp add: chain_reparam_weak_chain_def) have 0:" \x\one_chain1. (case x of (k, \) \ (real_of_int k * line_integral F basis \) = (case f x of (k, \) \ real_of_int k * line_integral F basis \) \ line_integral_exists F basis \)" proof {fix k1 \1 assume ass1: "(k1,\1) \one_chain1" have "real_of_int k1 * line_integral F basis \1 = (case (f (k1,\1)) of (k2, \2) \ real_of_int k2 * line_integral F basis \2) \ line_integral_exists F basis \1" proof(cases) assume ass2: "k1 = 1" let ?k2 = "fst (f (k1, \1))" let ?\2 = "snd (f (k1, \1))" have "real_of_int k1 * line_integral F basis \1 = real_of_int ?k2 * line_integral F basis ?\2\ line_integral_exists F basis \1" proof(cases) assume ass3: "?k2 = 1" then have 0: "reparam_weak \1 ?\2" using ass1 ass2 f(2) by auto have 1: "?\2 C1_differentiable_on {0..1}" using f(3) assms(2) ass1 by force have 2: "\b\basis. continuous_on (path_image ?\2) (\x. F x \ b)" using f(3) assms(3) ass1 by force show "real_of_int k1 * line_integral F basis \1 = real_of_int ?k2 * line_integral F basis ?\2 \ line_integral_exists F basis \1" using assms reparam_weak_eq_line_integrals_basis[OF 0 1 2 assms(4)] ass2 ass3 by auto next assume "?k2 \ 1" then have ass3: "?k2 = -1" using bound2 ass1 f(3) unfolding boundary_chain_def by force then have 0: "reparam_weak \1 (reversepath ?\2)" using ass1 ass2 f(2) by auto have 1: "(reversepath ?\2) C1_differentiable_on {0..1}" using f(3) assms(2) ass1 smooth_path_imp_reverse by force have 2: "\b\basis. continuous_on (path_image (reversepath ?\2)) (\x. F x \ b)" using f(3) assms(3) ass1 path_image_reversepath by force have 3: "line_integral F basis ?\2 = - line_integral F basis (reversepath ?\2)" proof- have i:"valid_path (reversepath ?\2) " using 1 C1_differentiable_imp_piecewise by (auto simp add: valid_path_def) show ?thesis using line_integral_on_reverse_path(1)[OF i line_integral_exists_smooth[OF 1 2 ]] assms by auto qed show "real_of_int k1 * line_integral F basis \1 = real_of_int ?k2 * line_integral F basis ?\2 \ line_integral_exists F basis \1" using assms reparam_weak_eq_line_integrals_basis[OF 0 1 2 assms(4)] ass2 ass3 3 by auto qed then show "real_of_int k1 * line_integral F basis \1 = (case f (k1, \1) of (k2, \2) \ real_of_int k2 * line_integral F basis \2) \ line_integral_exists F basis \1" by (simp add: case_prod_beta') next assume "k1 \ 1" then have ass2: "k1 = -1" using bound1 ass1 f(3) unfolding boundary_chain_def by force let ?k2 = "fst (f (k1, \1))" let ?\2 = "snd (f (k1, \1))" have "real_of_int k1 * line_integral F basis \1 = real_of_int ?k2 * line_integral F basis ?\2 \ line_integral_exists F basis \1" proof(cases) assume ass3: "?k2 = 1" then have 0: "reparam_weak \1 (reversepath ?\2)" using ass1 ass2 f(2) by auto have 1: "(reversepath ?\2) C1_differentiable_on {0..1}" using f(3) assms(2) ass1 smooth_path_imp_reverse by force have 2: "\b\basis. continuous_on (path_image (reversepath ?\2)) (\x. F x \ b)" using f(3) assms(3) ass1 path_image_reversepath by force have 3: "line_integral F basis ?\2 = - line_integral F basis (reversepath ?\2)" proof- have i:"valid_path (reversepath ?\2) " using 1 C1_differentiable_imp_piecewise by (auto simp add: valid_path_def) show ?thesis using line_integral_on_reverse_path(1)[OF i line_integral_exists_smooth[OF 1 2 assms(4)]] by auto qed show "real_of_int k1 * line_integral F basis \1 = real_of_int ?k2 * line_integral F basis ?\2 \ line_integral_exists F basis \1" using assms reparam_weak_eq_line_integrals_basis[OF 0 1 2 assms(4)] ass2 ass3 3 by auto next assume "?k2 \ 1" then have ass3: "?k2 = -1" using bound2 ass1 f(3) unfolding boundary_chain_def by force then have 0: "reparam_weak \1 ?\2" using ass1 ass2 f(2) by auto have 1: "?\2 C1_differentiable_on {0..1}" using f(3) assms(2) ass1 by force have 2: "\b\basis. continuous_on (path_image ?\2) (\x. F x \ b)" using f(3) assms(3) ass1 by force show "real_of_int k1 * line_integral F basis \1 = real_of_int ?k2 * line_integral F basis ?\2 \ line_integral_exists F basis \1" using assms reparam_weak_eq_line_integrals_basis[OF 0 1 2 assms(4)] ass2 ass3 by auto qed then show "real_of_int k1 * line_integral F basis \1 = (case f (k1, \1) of (k2, \2) \ real_of_int k2 * line_integral F basis \2) \ line_integral_exists F basis \1" by (simp add: case_prod_beta') qed } then show "\x. x \ one_chain1 \ (case x of (k, \) \ (real_of_int k * line_integral F basis \) = (case f x of (k, \) \ real_of_int k * line_integral F basis \) \ line_integral_exists F basis \)" by (auto simp add: case_prod_beta') qed show "one_chain_line_integral F basis one_chain1 = one_chain_line_integral F basis one_chain2" using 0 by(simp add: one_chain_line_integral_def sum_bij[OF f(1) _ f(3)] split_beta) show "\(k, \)\one_chain1. line_integral_exists F basis \" using 0 by blast qed definition chain_reparam_chain where "chain_reparam_chain one_chain1 one_chain2 \ \f. bij f \ f ` one_chain1 = one_chain2 \ (\(k,\)\one_chain1. if k = fst (f(k,\)) then reparam \ (snd (f(k,\))) else reparam \ (reversepath (snd (f(k,\)))))" definition chain_reparam_weak_path::"((real) \ (real * real)) \ ((int * ((real) \ (real * real))) set) \ bool" where "chain_reparam_weak_path \ one_chain \ \l. set l = one_chain \ distinct l \ reparam \ (rec_join l) \ valid_chain_list l \ l \ []" lemma chain_reparam_chain_line_integral: assumes "chain_reparam_chain one_chain1 one_chain2" "\(k2,\2)\one_chain2. \2 piecewise_C1_differentiable_on {0..1}" "\(k2,\2)\one_chain2.\b\basis. continuous_on (path_image \2) (\x. F x \ b)" "finite basis" and bound1: "boundary_chain one_chain1" and bound2: "boundary_chain one_chain2" and line: "\(k2,\2)\one_chain2. (\b\basis. line_integral_exists F {b} \2)" shows "one_chain_line_integral F basis one_chain1 = one_chain_line_integral F basis one_chain2" "\(k, \)\one_chain1. line_integral_exists F basis \" proof- obtain f where f: "bij f" "(\(k,\)\one_chain1. if k = fst (f(k,\)) then reparam \ (snd (f(k,\))) else reparam \ (reversepath (snd (f(k,\)))))" "f ` one_chain1 = one_chain2" using assms(1) by (auto simp add: chain_reparam_chain_def) have integ_exist_b: "\(k1,\1)\one_chain1. \b\basis. line_integral_exists F {b} (snd (f (k1, \1)))" using line f by fastforce have valid_cubes: "\(k1,\1)\one_chain1. valid_path (snd (f (k1, \1)))" using assms(2) f(3) valid_path_def by fastforce have integ_rev_exist_b: "\(k1,\1)\one_chain1. \b\basis. line_integral_exists F {b} (reversepath (snd (f (k1, \1))))" using line_integral_on_reverse_path(2) integ_exist_b valid_cubes by blast have 0:" \x\one_chain1. (case x of (k, \) \ (real_of_int k * line_integral F basis \) = (case f x of (k, \) \ real_of_int k * line_integral F basis \) \ line_integral_exists F basis \)" proof {fix k1 \1 assume ass1: "(k1,\1) \one_chain1" have "real_of_int k1 * line_integral F basis \1 = (case (f (k1,\1)) of (k2, \2) \ real_of_int k2 * line_integral F basis \2) \ line_integral_exists F basis \1" proof(cases) assume ass2: "k1 = 1" let ?k2 = "fst (f (k1, \1))" let ?\2 = "snd (f (k1, \1))" have "real_of_int k1 * line_integral F basis \1 = real_of_int ?k2 * line_integral F basis ?\2\ line_integral_exists F basis \1" proof(cases) assume ass3: "?k2 = 1" then have 0: "reparam \1 ?\2" using ass1 ass2 f(2) by auto have 1: "?\2 piecewise_C1_differentiable_on {0..1}" using f(3) assms(2) ass1 by force have 2: "\b\basis. continuous_on (path_image ?\2) (\x. F x \ b)" using f(3) assms(3) ass1 by force show "real_of_int k1 * line_integral F basis \1 = real_of_int ?k2 * line_integral F basis ?\2 \ line_integral_exists F basis \1" using assms reparam_eq_line_integrals_basis[OF 0 1 2 assms(4)] integ_exist_b ass1 ass2 ass3 by auto next assume "?k2 \ 1" then have ass3: "?k2 = -1" using bound2 ass1 f(3) unfolding boundary_chain_def by force then have 0: "reparam \1 (reversepath ?\2)" using ass1 ass2 f(2) by auto have 1: "(reversepath ?\2) piecewise_C1_differentiable_on {0..1}" using f(3) assms(2) ass1 piecewise_smooth_path_imp_reverse by force have 2: "\b\basis. continuous_on (path_image (reversepath ?\2)) (\x. F x \ b)" using f(3) assms(3) ass1 path_image_reversepath by force have 3: "line_integral F basis ?\2 = - line_integral F basis (reversepath ?\2)" proof- have i:"valid_path (reversepath ?\2) " using 1 C1_differentiable_imp_piecewise by (auto simp add: valid_path_def) have ii: "line_integral_exists F basis (snd (f (k1, \1)))" using assms(4) line_integral_sum_basis(2) integ_exist_b ass1 by fastforce show ?thesis using i ii line_integral_on_reverse_path(1) valid_path_reversepath by blast qed show "real_of_int k1 * line_integral F basis \1 = real_of_int ?k2 * line_integral F basis ?\2 \ line_integral_exists F basis \1" using assms reparam_eq_line_integrals_basis[OF 0 1 2 assms(4)] integ_rev_exist_b ass1 ass2 ass3 3 by auto qed then show "real_of_int k1 * line_integral F basis \1 = (case f (k1, \1) of (k2, \2) \ real_of_int k2 * line_integral F basis \2) \ line_integral_exists F basis \1" by (simp add: case_prod_beta') next assume "k1 \ 1" then have ass2: "k1 = -1" using bound1 ass1 f(3) unfolding boundary_chain_def by force let ?k2 = "fst (f (k1, \1))" let ?\2 = "snd (f (k1, \1))" have "real_of_int k1 * line_integral F basis \1 = real_of_int ?k2 * line_integral F basis ?\2 \ line_integral_exists F basis \1" proof(cases) assume ass3: "?k2 = 1" then have 0: "reparam \1 (reversepath ?\2)" using ass1 ass2 f(2) by auto have 1: "(reversepath ?\2) piecewise_C1_differentiable_on {0..1}" using f(3) assms(2) ass1 piecewise_smooth_path_imp_reverse by force have 2: "\b\basis. continuous_on (path_image (reversepath ?\2)) (\x. F x \ b)" using f(3) assms(3) ass1 path_image_reversepath by force have 3: "line_integral F basis ?\2 = - line_integral F basis (reversepath ?\2)" proof- have i:"valid_path (reversepath ?\2) " using 1 C1_differentiable_imp_piecewise by (auto simp add: valid_path_def) show ?thesis using line_integral_on_reverse_path(1)[OF i] integ_rev_exist_b using ass1 assms(4) line_integral_sum_basis(2) by fastforce qed show "real_of_int k1 * line_integral F basis \1 = real_of_int ?k2 * line_integral F basis ?\2 \ line_integral_exists F basis \1" using assms reparam_eq_line_integrals_basis[OF 0 1 2 assms(4)] ass2 ass3 3 using ass1 integ_rev_exist_b by auto next assume "?k2 \ 1" then have ass3: "?k2 = -1" using bound2 ass1 f(3) unfolding boundary_chain_def by force then have 0: "reparam \1 ?\2" using ass1 ass2 f(2) by auto have 1: "?\2 piecewise_C1_differentiable_on {0..1}" using f(3) assms(2) ass1 by force have 2: "\b\basis. continuous_on (path_image ?\2) (\x. F x \ b)" using f(3) assms(3) ass1 by force show "real_of_int k1 * line_integral F basis \1 = real_of_int ?k2 * line_integral F basis ?\2 \ line_integral_exists F basis \1" using assms reparam_eq_line_integrals_basis[OF 0 1 2 assms(4)] ass2 ass3 using ass1 integ_exist_b by auto qed then show "real_of_int k1 * line_integral F basis \1 = (case f (k1, \1) of (k2, \2) \ real_of_int k2 * line_integral F basis \2) \ line_integral_exists F basis \1" by (simp add: case_prod_beta') qed } then show "\x. x \ one_chain1 \ (case x of (k, \) \ (real_of_int k * line_integral F basis \) = (case f x of (k, \) \ real_of_int k * line_integral F basis \) \ line_integral_exists F basis \)" by (auto simp add: case_prod_beta') qed show "one_chain_line_integral F basis one_chain1 = one_chain_line_integral F basis one_chain2" using 0 by (simp add: one_chain_line_integral_def sum_bij[OF f(1) _ f(3)] prod.case_eq_if) show "\(k, \)\one_chain1. line_integral_exists F basis \" using 0 by blast qed lemma path_image_rec_join: fixes \::"real \ (real \ real)" fixes k::int fixes l shows "\k \. (k, \) \ set l \ valid_chain_list l \ path_image \ \ path_image (rec_join l)" proof(induction l) case Nil then show ?case by auto next case ass: (Cons a l) obtain k' \' where a: "a = (k',\')" by force have "path_image \ \ path_image (rec_join ((k',\') # l))" proof(cases) assume "l=[]" then show "path_image \ \ path_image (rec_join ((k',\') # l))" using ass(2-3) a by(auto simp add:) next assume "l\[]" then obtain b l' where b_l: "l = b # l'" by (meson rec_join.elims) obtain k'' \'' where b: "b = (k'',\'')" by force show ?thesis using ass path_image_reversepath b_l path_image_join by(fastforce simp add: a) qed then show ?case using a by auto qed lemma path_image_rec_join_2: fixes l shows "l \ [] \ valid_chain_list l \ path_image (rec_join l) \ (\(k, \) \ set l. path_image \)" proof(induction l) case Nil then show ?case by auto next case ass: (Cons a l) obtain k' \' where a: "a = (k',\')" by force have "path_image (rec_join (a # l)) \ (\(k, y)\set (a # l). path_image y)" proof(cases) assume "l=[]" then show "path_image (rec_join (a # l)) \ (\(k, y)\set (a # l). path_image y)" using step a by(auto simp add:) next assume "l\[]" then obtain b l' where b_l: "l = b # l'" by (meson rec_join.elims) obtain k'' \'' where b: "b = (k'',\'')" by force show ?thesis using ass path_image_reversepath b_l path_image_join apply(auto simp add: a) (*Excuse the ugliness of the next script*) apply blast by fastforce qed then show ?case using a by auto qed lemma continuous_on_closed_UN: assumes "finite S" shows "((\s. s \ S \ closed s) \ (\s. s \ S \ continuous_on s f) \ continuous_on (\S) f)" using assms proof(induction S) case empty then show ?case by auto next case (insert x F) then show ?case using continuous_on_closed_Un closed_Union by (simp add: closed_Union continuous_on_closed_Un) qed lemma chain_reparam_weak_path_line_integral: assumes path_eq_chain: "chain_reparam_weak_path \ one_chain" and boundary_chain: "boundary_chain one_chain" and line_integral_exists: "\b\basis. \(k::int, \) \ one_chain. line_integral_exists F {b} \" and valid_path: "\(k::int, \) \ one_chain. valid_path \" and finite_basis: "finite basis" and cont: "\b\basis. \(k,\2)\one_chain. continuous_on (path_image \2) (\x. F x \ b)" and finite_one_chain: "finite one_chain" shows "line_integral F basis \ = one_chain_line_integral F basis one_chain" "line_integral_exists F basis \" (*"valid_path \" This cannot be proven, see the crappy assumption of derivs/eq_pw_smooth :( *) proof - obtain l where l_props: "set l = one_chain" "distinct l" "reparam \ (rec_join l)" "valid_chain_list l" "l \ []" using chain_reparam_weak_path_def assms by auto have bchain_l: "boundary_chain (set l)" using l_props boundary_chain by (simp add: boundary_chain_def) have cont_forall: "\b\basis. continuous_on (\(k, \)\one_chain. path_image \) (\x. F x \ b)" proof fix b assume ass: "b \ basis" have "continuous_on (\((path_image \ snd) ` one_chain)) (\x. F x \ b)" apply(rule continuous_on_closed_UN[where ?S = "(path_image o snd) ` one_chain" ]) proof show "finite one_chain" using finite_one_chain by auto show "\s. s \ (path_image \ snd) ` one_chain \ closed s" using closed_valid_path_image[OF] valid_path by fastforce show "\s. s \ (path_image \ snd) ` one_chain \ continuous_on s (\x. F x \ b)" using cont ass by force qed then show "continuous_on (\(k, \)\one_chain. path_image \) (\x. F x \ b)" by (auto simp add: Union_eq case_prod_beta) qed show "line_integral_exists F basis \" proof (rule reparam_eq_line_integrals_basis[OF l_props(3) _ _ finite_basis]) let ?\2.0="rec_join l" show "?\2.0 piecewise_C1_differentiable_on {0..1}" apply(simp only: valid_path_def[symmetric]) apply(rule joined_is_valid) using assms l_props by auto show "\b\basis. continuous_on (path_image (rec_join l)) (\x. F x \ b)" using path_image_rec_join_2[OF l_props(5) l_props(4)] l_props(1) using cont_forall continuous_on_subset by blast show "\b\basis. line_integral_exists F {b} (rec_join l)" proof fix b assume ass: "b\basis" show "line_integral_exists F {b} (rec_join l)" proof (rule line_integral_exists_on_rec_join) show "boundary_chain (set l)" using l_props boundary_chain by auto show "valid_chain_list l" using l_props by auto show "\k \. (k, \) \ set l \ valid_path \" using l_props assms by auto show "\(k, \)\set l. line_integral_exists F {b} \" using l_props line_integral_exists ass by blast qed qed qed show "line_integral F basis \ = one_chain_line_integral F basis one_chain" proof- have i: "line_integral F basis (rec_join l) = one_chain_line_integral F basis one_chain" proof (rule one_chain_line_integral_rec_join) show "set l = one_chain" "distinct l" "valid_chain_list l" using l_props by auto show "boundary_chain one_chain" using boundary_chain by auto show "\(k, \)\one_chain. line_integral_exists F basis \" using line_integral_sum_basis(2)[OF finite_basis] line_integral_exists by blast show "\(k, \)\one_chain. valid_path \" using valid_path by auto show "finite basis" using finite_basis by auto qed have ii: "line_integral F basis \ = line_integral F basis (rec_join l)" proof (rule reparam_eq_line_integrals_basis[OF l_props(3) _ _ finite_basis]) let ?\2.0="rec_join l" show "?\2.0 piecewise_C1_differentiable_on {0..1}" apply(simp only: valid_path_def[symmetric]) apply(rule joined_is_valid) using assms l_props by auto show "\b\basis. continuous_on (path_image (rec_join l)) (\x. F x \ b)" using path_image_rec_join_2[OF l_props(5) l_props(4)] l_props(1) using cont_forall continuous_on_subset by blast show "\b\basis. line_integral_exists F {b} (rec_join l)" proof fix b assume ass: "b\basis" show "line_integral_exists F {b} (rec_join l)" proof (rule line_integral_exists_on_rec_join) show "boundary_chain (set l)" using l_props boundary_chain by auto show "valid_chain_list l" using l_props by auto show "\k \. (k, \) \ set l \ valid_path \" using l_props assms by auto show "\(k, \)\set l. line_integral_exists F {b} \" using l_props line_integral_exists ass by blast qed qed qed show "line_integral F basis \ = one_chain_line_integral F basis one_chain" using i ii by auto qed qed definition chain_reparam_chain' where "chain_reparam_chain' one_chain1 subdiv \ \f. ((\(f ` one_chain1)) = subdiv) \ (\cube \ one_chain1. chain_reparam_weak_path (rec_join [cube]) (f cube)) \ (\p\one_chain1. \p'\one_chain1. p \ p' \ f p \ f p' = {}) \ (\x \ one_chain1. finite (f x))" lemma chain_reparam_chain'_imp_finite_subdiv: assumes "finite one_chain1" "chain_reparam_chain' one_chain1 subdiv" shows "finite subdiv" using assms by(auto simp add: chain_reparam_chain'_def) lemma chain_reparam_chain'_line_integral: assumes chain1_eq_chain2: "chain_reparam_chain' one_chain1 subdiv" and boundary_chain1: "boundary_chain one_chain1" and boundary_chain2: "boundary_chain subdiv" and line_integral_exists_on_chain2: "\b\basis. \(k::int, \) \ subdiv. line_integral_exists F {b} \" and valid_path: "\(k, \) \ subdiv. valid_path \" and valid_path_2: "\(k, \) \ one_chain1. valid_path \" and finite_chain1: "finite one_chain1" and finite_basis: "finite basis" and cont_field: " \b\basis. \(k, \2)\subdiv. continuous_on (path_image \2) (\x. F x \ b)" shows "one_chain_line_integral F basis one_chain1 = one_chain_line_integral F basis subdiv" "\(k, \) \ one_chain1. line_integral_exists F basis \" proof - obtain f where f_props: "((\(f ` one_chain1)) = subdiv)" "(\cube \ one_chain1. chain_reparam_weak_path (rec_join [cube]) (f cube))" "(\p\one_chain1. \p'\one_chain1. p \ p' \ f p \ f p' = {})" "(\x \ one_chain1. finite (f x))" using chain1_eq_chain2 by (auto simp add: chain_reparam_chain'_def) then have 0: "one_chain_line_integral F basis subdiv = one_chain_line_integral F basis (\(f ` one_chain1))" by auto {fix k \ assume ass:"(k, \) \ one_chain1" have "line_integral_exists F basis \ \ one_chain_line_integral F basis (f (k, \)) = k * line_integral F basis \" proof(cases "k = 1") assume k_eq_1: "k = 1" then have i:"chain_reparam_weak_path \ (f(k,\))" using f_props(2) ass by auto have ii:"boundary_chain (f(k,\))" using f_props(1) ass assms unfolding boundary_chain_def by blast have iii:"\b\basis. \(k, \)\f (k, \). line_integral_exists F {b} \" using f_props(1) ass assms by blast have "iv": " \(k, \)\f (k, \). valid_path \" using f_props(1) ass assms by blast have v: "\b\basis. \(k, \2)\f (k, \). continuous_on (path_image \2) (\x. F x \ b)" using f_props(1) ass assms by blast have "line_integral_exists F basis \ \ one_chain_line_integral F basis (f (k, \)) = line_integral F basis \" using chain_reparam_weak_path_line_integral[OF i ii iii iv finite_basis v] ass f_props(4) by (auto) then show "line_integral_exists F basis \ \ one_chain_line_integral F basis (f (k, \)) = k * line_integral F basis \" using k_eq_1 by auto next assume "k \ 1" then have k_eq_neg1: "k = -1" using ass boundary_chain1 by (auto simp add: boundary_chain_def) then have i:"chain_reparam_weak_path (reversepath \) (f(k,\))" using f_props(2) ass by auto have ii:"boundary_chain (f(k,\))" using f_props(1) ass assms unfolding boundary_chain_def by blast have iii:"\b\basis. \(k, \)\f (k, \). line_integral_exists F {b} \" using f_props(1) ass assms by blast have "iv": " \(k, \)\f (k, \). valid_path \" using f_props(1) ass assms by blast have v: "\b\basis. \(k, \2)\f (k, \). continuous_on (path_image \2) (\x. F x \ b)" using f_props(1) ass assms by blast have x:"line_integral_exists F basis (reversepath \) \ one_chain_line_integral F basis (f (k, \)) = line_integral F basis (reversepath \)" using chain_reparam_weak_path_line_integral[OF i ii iii iv finite_basis v] ass f_props(4) by auto have "valid_path (reversepath \)" using f_props(1) ass assms by auto then have "line_integral_exists F basis \ \ one_chain_line_integral F basis (f (k, \)) = - (line_integral F basis \)" using line_integral_on_reverse_path reversepath_reversepath x ass by metis then show "line_integral_exists F basis \ \ one_chain_line_integral F basis (f (k::int, \)) = k * line_integral F basis \" using k_eq_neg1 by auto qed} note cube_line_integ = this have finite_chain2: "finite subdiv" using finite_chain1 f_props(1) f_props(4) by auto show line_integral_exists_on_chain1: "\(k, \) \ one_chain1. line_integral_exists F basis \" using cube_line_integ by auto have 1: "one_chain_line_integral F basis (\(f ` one_chain1)) = one_chain_line_integral F basis one_chain1" proof - have 0:"one_chain_line_integral F basis (\(f ` one_chain1)) = (\one_chain \ (f ` one_chain1). one_chain_line_integral F basis one_chain)" proof - have finite: "\chain \ (f ` one_chain1). finite chain" using f_props(1) finite_chain2 by (meson Sup_upper finite_subset) have disj: " \A\f ` one_chain1. \B\f ` one_chain1. A \ B \ A \ B = {}" apply(simp add:image_def) using f_props(3) by metis show "one_chain_line_integral F basis (\(f ` one_chain1)) = (\one_chain \ (f ` one_chain1). one_chain_line_integral F basis one_chain)" using Groups_Big.comm_monoid_add_class.sum.Union_disjoint[OF finite disj] one_chain_line_integral_def by auto qed have 1:"(\one_chain \ (f ` one_chain1). one_chain_line_integral F basis one_chain) = one_chain_line_integral F basis one_chain1" proof - have "(\one_chain \ (f ` one_chain1). one_chain_line_integral F basis one_chain) = (\(k,\)\one_chain1. k*(line_integral F basis \))" proof - have i:"(\one_chain \ (f ` (one_chain1 - {p. f p = {}})). one_chain_line_integral F basis one_chain) = (\(k,\)\one_chain1 - {p. f p = {}}. k*(line_integral F basis \))" proof - have "inj_on f (one_chain1 - {p. f p = {}})" unfolding inj_on_def using f_props(3) by force then have 0: "(\one_chain \ (f ` (one_chain1 - {p. f p = {}})). one_chain_line_integral F basis one_chain) = (\(k, \) \ (one_chain1 - {p. f p = {}}). one_chain_line_integral F basis (f (k, \)))" using Groups_Big.comm_monoid_add_class.sum.reindex by auto have "\ k \. (k, \) \ (one_chain1 - {p. f p = {}}) \ one_chain_line_integral F basis (f(k, \)) = k* (line_integral F basis \)" using cube_line_integ by auto then have "(\(k, \) \ (one_chain1 - {p. f p = {}}). one_chain_line_integral F basis (f (k, \))) = (\(k, \) \ (one_chain1 - {p. f p = {}}). k* (line_integral F basis \))" by (auto intro!: Finite_Cartesian_Product.sum_cong_aux) then show "(\one_chain \ (f ` (one_chain1 - {p. f p = {}})). one_chain_line_integral F basis one_chain) = (\(k, \) \ (one_chain1 - {p. f p = {}}). k* (line_integral F basis \))" using 0 by auto qed have "\ (k::int) \. (k, \) \ one_chain1 \ (f (k, \) = {}) \ (k, \) \ {(k',\'). k'* (line_integral F basis \') = 0}" proof- fix k::"int" fix \::"real\real*real" assume ass:"(k, \) \ one_chain1" "(f (k, \) = {})" then have zero_line_integral:"one_chain_line_integral F basis (f (k, \)) = 0" using one_chain_line_integral_def by auto show "(k, \) \ {(k'::int, \'). k' * line_integral F basis \' = 0}" using zero_line_integral cube_line_integ ass by force qed then have ii:"(\one_chain \ (f ` (one_chain1 - {p. f p = {}})). one_chain_line_integral F basis one_chain) = (\one_chain \ (f ` (one_chain1)). one_chain_line_integral F basis one_chain)" proof - have "\one_chain. one_chain \ (f ` (one_chain1)) - (f ` (one_chain1 - {p. f p = {}})) \ one_chain_line_integral F basis one_chain = 0" proof - fix one_chain assume "one_chain \ (f ` (one_chain1)) - (f ` (one_chain1 - {p. f p = {}}))" then have "one_chain = {}" by auto then show "one_chain_line_integral F basis one_chain = 0" by (auto simp add: one_chain_line_integral_def) qed then have 0:"(\one_chain \ f ` (one_chain1) - (f ` (one_chain1 - {p. f p = {}})). one_chain_line_integral F basis one_chain) = 0" using Groups_Big.comm_monoid_add_class.sum.neutral by auto then have "(\one_chain \ f ` (one_chain1). one_chain_line_integral F basis one_chain) - (\one_chain \ (f ` (one_chain1 - {p. f p = {}})). one_chain_line_integral F basis one_chain) = 0" proof - have finte: "finite (f ` one_chain1)" using finite_chain1 by auto show ?thesis using Groups_Big.sum_diff[OF finte, of " (f ` (one_chain1 - {p. f p = {}}))" "one_chain_line_integral F basis"] 0 by auto qed then show "(\one_chain \ (f ` (one_chain1 - {p. f p = {}})). one_chain_line_integral F basis one_chain) = (\one_chain \ (f ` (one_chain1)). one_chain_line_integral F basis one_chain)" by auto qed have "\ (k::int) \. (k, \) \ one_chain1 \ (f (k, \) = {}) \ f (k, \) \ {chain. one_chain_line_integral F basis chain = 0}" proof- fix k::"int" fix \::"real\real*real" assume ass:"(k, \) \ one_chain1" "(f (k, \) = {})" then have "one_chain_line_integral F basis (f (k, \)) = 0" using one_chain_line_integral_def by auto then show "f (k, \) \ {p'. (one_chain_line_integral F basis p' = 0)}" by auto qed then have iii:"(\(k::int,\)\one_chain1 - {p. f p = {}}. k*(line_integral F basis \)) = (\(k::int,\)\one_chain1. k*(line_integral F basis \))" proof- have "\ k \. (k,\)\one_chain1 - (one_chain1 - {p. f p = {}}) \ k* (line_integral F basis \) = 0" proof- fix k \ assume ass: "(k,\)\one_chain1 - (one_chain1 - {p. f p = {}})" then have "f(k, \) = {}" by auto then have "one_chain_line_integral F basis (f(k, \)) = 0" by (auto simp add: one_chain_line_integral_def) then have zero_line_integral:"one_chain_line_integral F basis (f (k, \)) = 0" using one_chain_line_integral_def by auto then show "k* (line_integral F basis \) = 0" using zero_line_integral cube_line_integ ass by force qed then have "\(k::int,\)\one_chain1 - (one_chain1 - {p. f p = {}}). k* (line_integral F basis \) = 0" by auto then have "(\(k::int,\)\one_chain1 - (one_chain1 - {p. f p = {}}). k*(line_integral F basis \)) = 0" using Groups_Big.comm_monoid_add_class.sum.neutral [of "one_chain1 - (one_chain1 - {p. f p = {}})" "(\(k::int,\). k* (line_integral F basis \))"] by (simp add: split_beta) then have "(\(k::int,\)\one_chain1. k*(line_integral F basis \)) - (\(k::int,\)\ (one_chain1 - {p. f p = {}}). k*(line_integral F basis \)) = 0" using Groups_Big.sum_diff[OF finite_chain1] by (metis (no_types) Diff_subset \(\(k, \)\one_chain1 - (one_chain1 - {p. f p = {}}). k * line_integral F basis \) = 0\ \\f B. B \ one_chain1 \ sum f (one_chain1 - B) = sum f one_chain1 - sum f B\) then show ?thesis by auto qed show ?thesis using i ii iii by auto qed then show ?thesis using one_chain_line_integral_def by auto qed show ?thesis using 0 1 by auto qed show "one_chain_line_integral F basis one_chain1 = one_chain_line_integral F basis subdiv" using 0 1 by auto qed lemma chain_reparam_chain'_line_integral_smooth_cubes: assumes "chain_reparam_chain' one_chain1 one_chain2" "\(k2,\2)\one_chain2. \2 C1_differentiable_on {0..1}" "\b\basis.\(k2,\2)\one_chain2. continuous_on (path_image \2) (\x. F x \ b)" "finite basis" "finite one_chain1" "boundary_chain one_chain1" "boundary_chain one_chain2" "\(k,\)\one_chain1. valid_path \" shows "one_chain_line_integral F basis one_chain1 = one_chain_line_integral F basis one_chain2" "\(k, \)\one_chain1. line_integral_exists F basis \" proof- {fix b assume "b \ basis" fix k \ assume "(k, \)\one_chain2" have "line_integral_exists F {b} \" apply(rule line_integral_exists_smooth) using \(k, \) \ one_chain2\ assms(2) apply blast using assms using \(k, \) \ one_chain2\ \b \ basis\ apply blast using \b \ basis\ by blast} then have a: "\b\basis. \(k, \)\one_chain2. line_integral_exists F {b} \" by auto have b: "\(k2,\2)\one_chain2. valid_path \2" using assms(2) by (simp add: C1_differentiable_imp_piecewise case_prod_beta valid_path_def) show "one_chain_line_integral F basis one_chain1 = one_chain_line_integral F basis one_chain2" by (rule chain_reparam_chain'_line_integral[OF assms(1) assms(6) assms(7) a b assms(8) assms(5) assms(4) assms(3)]) show "\(k, \)\one_chain1. line_integral_exists F basis \" by (rule chain_reparam_chain'_line_integral[OF assms(1) assms(6) assms(7) a b assms(8) assms(5) assms(4) assms(3)]) qed lemma chain_subdiv_path_pathimg_subset: assumes "chain_subdiv_path \ subdiv" shows "\(k',\')\subdiv. (path_image \') \ path_image \" using assms chain_subdiv_path.simps path_image_rec_join by(auto simp add: chain_subdiv_path.simps) lemma reparam_path_image: assumes "reparam \1 \2" shows "path_image \1 = path_image \2" using assms apply (auto simp add: reparam_def path_image_def image_def bij_betw_def) apply (force dest!: equalityD2) done lemma chain_reparam_weak_path_pathimg_subset: assumes "chain_reparam_weak_path \ subdiv" shows "\(k',\')\subdiv. (path_image \') \ path_image \" using assms apply(auto simp add: chain_reparam_weak_path_def) using path_image_rec_join reparam_path_image by blast lemma chain_subdiv_chain_pathimg_subset': assumes "chain_subdiv_chain one_chain subdiv" assumes "(k,\)\ subdiv" shows " \k' \'. (k',\')\ one_chain \ path_image \ \ path_image \'" using assms unfolding chain_subdiv_chain_def pairwise_def apply auto by (metis chain_subdiv_path.cases coeff_cube_to_path.simps path_image_rec_join path_image_reversepath) lemma chain_subdiv_chain_pathimg_subset: assumes "chain_subdiv_chain one_chain subdiv" shows "\ (path_image ` {\. \k. (k,\)\ subdiv}) \ \ (path_image ` {\. \k. (k,\)\ one_chain})" using assms unfolding chain_subdiv_chain_def pairwise_def apply auto by (metis UN_iff assms chain_subdiv_chain_pathimg_subset' subsetCE case_prodI2) lemma chain_reparam_chain'_pathimg_subset': assumes "chain_reparam_chain' one_chain subdiv" assumes "(k,\)\ subdiv" shows " \k' \'. (k',\')\ one_chain \ path_image \ \ path_image \'" using assms chain_reparam_weak_path_pathimg_subset apply(auto simp add: chain_reparam_chain'_def set_eq_iff) using path_image_reversepath case_prodE case_prodD old.prod.exhaust apply (simp add: list.distinct(1) list.inject rec_join.elims) by (smt case_prodD coeff_cube_to_path.simps rec_join.simps(2) reversepath_simps(2) surj_pair) definition common_reparam_exists:: "(int \ (real \ real \ real)) set \ (int \ (real \ real \ real)) set \ bool" where "common_reparam_exists one_chain1 one_chain2 \ (\subdiv ps1 ps2. chain_reparam_chain' (one_chain1 - ps1) subdiv \ chain_reparam_chain' (one_chain2 - ps2) subdiv \ (\(k, \)\subdiv. \ C1_differentiable_on {0..1}) \ boundary_chain subdiv \ (\(k, \)\ps1. point_path \) \ (\(k, \)\ps2. point_path \))" lemma common_reparam_exists_imp_eq_line_integral: assumes finite_basis: "finite basis" and "finite one_chain1" "finite one_chain2" "boundary_chain (one_chain1::(int \ (real \ real \ real)) set)" "boundary_chain (one_chain2::(int \ (real \ real \ real)) set)" " \(k2, \2)\one_chain2. \b\basis. continuous_on (path_image \2) (\x. F x \ b)" "(common_reparam_exists one_chain1 one_chain2)" "(\(k, \)\one_chain1. valid_path \)" "(\(k, \)\one_chain2. valid_path \)" shows "one_chain_line_integral F basis one_chain1 = one_chain_line_integral F basis one_chain2" "\(k, \)\one_chain1. line_integral_exists F basis \" proof- obtain subdiv ps1 ps2 where props: "chain_reparam_chain' (one_chain1 - ps1) subdiv" "chain_reparam_chain' (one_chain2 - ps2) subdiv " "(\(k, \)\subdiv. \ C1_differentiable_on {0..1})" "boundary_chain subdiv" "(\(k, \)\ps1. point_path \)" "(\(k, \)\ps2. point_path \)" using assms by (auto simp add: common_reparam_exists_def) have subdiv_valid: "(\(k, \)\subdiv. valid_path \)" apply(simp add: valid_path_def) using props(3) using C1_differentiable_imp_piecewise by blast have onechain_boundary1: "boundary_chain (one_chain1 - ps1)" using assms(4) by (auto simp add: boundary_chain_def) have onechain_boundary2: "boundary_chain (one_chain2 - ps1)" using assms(5) by (auto simp add: boundary_chain_def) {fix k2 \2 b assume ass: "(k2, \2)\subdiv "" b\basis" have "\ k \. (k, \) \ subdiv \ \k' \'. (k', \') \ one_chain2 \ path_image \ \ path_image \'" by (meson chain_reparam_chain'_pathimg_subset' props Diff_subset subsetCE) then have "continuous_on (path_image \2) (\x. F x \ b)" using assms(6) continuous_on_subset[where ?f = " (\x. F x \ b)"] ass apply(auto simp add: subset_iff) by (metis (mono_tags, lifting) case_prodD)} then have cont_field: "\b\basis. \(k2, \2)\subdiv. continuous_on (path_image \2) (\x. F x \ b)" by auto have one_chain1_ps_valid: "(\(k, \)\one_chain1 - ps1. valid_path \)" using assms by auto have one_chain2_ps_valid: "(\(k, \)\one_chain2 - ps1. valid_path \)" using assms by auto have 0: "one_chain_line_integral F basis (one_chain1 - ps1) = one_chain_line_integral F basis subdiv" apply(rule chain_reparam_chain'_line_integral_smooth_cubes[OF props(1) props(3) cont_field finite_basis]) using props assms apply blast using props assms using onechain_boundary1 apply blast using props assms apply blast using one_chain1_ps_valid by blast have 1:"\(k, \)\(one_chain1 - ps1). line_integral_exists F basis \" apply(rule chain_reparam_chain'_line_integral_smooth_cubes[OF props(1) props(3) cont_field finite_basis]) using props assms apply blast using props assms using onechain_boundary1 apply blast using props assms apply blast using one_chain1_ps_valid by blast have 2: "one_chain_line_integral F basis (one_chain2 - ps2) = one_chain_line_integral F basis subdiv" apply(rule chain_reparam_chain'_line_integral_smooth_cubes[OF props(2) props(3) cont_field finite_basis]) using props assms apply blast apply (simp add: assms(5) boundary_chain_diff) apply (simp add: props(4)) by (simp add: assms(9)) have 3:"\(k, \)\(one_chain2 - ps2). line_integral_exists F basis \" apply(rule chain_reparam_chain'_line_integral_smooth_cubes[OF props(2) props(3) cont_field finite_basis]) using props assms apply blast apply (simp add: assms(5) boundary_chain_diff) apply (simp add: props(4)) by (simp add: assms(9)) show line_int_ex_chain1: "\(k, \)\one_chain1. line_integral_exists F basis \" using 0 using "1" finite_basis line_integral_exists_point_path props(5) by fastforce then show "one_chain_line_integral F basis one_chain1 = one_chain_line_integral F basis one_chain2" using 0 1 2 3 using assms(2-3) finite_basis one_chain_line_integral_point_paths props(5) props(6) by auto qed definition subcube :: "real \ real \ (int \ (real \ real \ real)) \ (int \ (real \ real \ real))" where "subcube a b cube = (fst cube, subpath a b (snd cube))" lemma subcube_valid_path: assumes "valid_path (snd cube)" "a \ {0..1}" "b \ {0..1}" shows "valid_path (snd (subcube a b cube))" using valid_path_subpath[OF assms] by (auto simp add: subcube_def) end diff --git a/thys/Hybrid_Systems_VCs/HS_ODEs.thy b/thys/Hybrid_Systems_VCs/HS_ODEs.thy --- a/thys/Hybrid_Systems_VCs/HS_ODEs.thy +++ b/thys/Hybrid_Systems_VCs/HS_ODEs.thy @@ -1,571 +1,571 @@ (* Title: ODEs and Dynamical Systems for HS verification Author: Jonathan Julián Huerta y Munive, 2019 Maintainer: Jonathan Julián Huerta y Munive *) section \ Ordinary Differential Equations \ text \Vector fields @{text "f::real \ 'a \ ('a::real_normed_vector)"} represent systems of ordinary differential equations (ODEs). Picard-Lindeloef's theorem guarantees existence and uniqueness of local solutions to initial value problems involving Lipschitz continuous vector fields. A (local) flow @{text "\::real \ 'a \ ('a::real_normed_vector)"} for such a system is the function that maps initial conditions to their unique solutions. In dynamical systems, the set of all points @{text "\ t s::'a"} for a fixed @{text "s::'a"} is the flow's orbit. If the orbit of each @{text "s \ I"} is conatined in @{text I}, then @{text I} is an invariant set of this system. This section formalises these concepts with a focus on hybrid systems (HS) verification.\ theory HS_ODEs imports "HS_Preliminaries" begin subsection \ Initial value problems and orbits \ notation image ("\

") lemma image_le_pred[simp]: "(\

f A \ {s. G s}) = (\x\A. G (f x))" unfolding image_def by force definition ivp_sols :: "(real \ 'a \ ('a::real_normed_vector)) \ real set \ 'a set \ real \ 'a \ (real \ 'a) set" ("Sols") where "Sols f T S t\<^sub>0 s = {X |X. (D X = (\t. f t (X t)) on T) \ X t\<^sub>0 = s \ X \ T \ S}" lemma ivp_solsI: assumes "D X = (\t. f t (X t)) on T" "X t\<^sub>0 = s" "X \ T \ S" shows "X \ Sols f T S t\<^sub>0 s" using assms unfolding ivp_sols_def by blast lemma ivp_solsD: assumes "X \ Sols f T S t\<^sub>0 s" shows "D X = (\t. f t (X t)) on T" and "X t\<^sub>0 = s" and "X \ T \ S" using assms unfolding ivp_sols_def by auto abbreviation "down T t \ {\\T. \\ t}" definition g_orbit :: "(('a::ord) \ 'b) \ ('b \ bool) \ 'a set \ 'b set" ("\") where "\ X G T = \{\

X (down T t) |t. \

X (down T t) \ {s. G s}}" lemma g_orbit_eq: fixes X::"('a::preorder) \ 'b" shows "\ X G T = {X t |t. t \ T \ (\\\down T t. G (X \))}" unfolding g_orbit_def apply safe using le_left_mono by blast auto definition g_orbital :: "('a \ 'a) \ ('a \ bool) \ real set \ 'a set \ real \ ('a::real_normed_vector) \ 'a set" where "g_orbital f G T S t\<^sub>0 s = \{\ X G T |X. X \ ivp_sols (\t. f) T S t\<^sub>0 s}" lemma g_orbital_eq: "g_orbital f G T S t\<^sub>0 s = {X t |t X. t \ T \ \

X (down T t) \ {s. G s} \ X \ Sols (\t. f) T S t\<^sub>0 s }" unfolding g_orbital_def ivp_sols_def g_orbit_eq image_le_pred by auto lemma g_orbitalI: assumes "X \ Sols (\t. f) T S t\<^sub>0 s" and "t \ T" and "(\

X (down T t) \ {s. G s})" shows "X t \ g_orbital f G T S t\<^sub>0 s" using assms unfolding g_orbital_eq(1) by auto lemma g_orbitalD: assumes "s' \ g_orbital f G T S t\<^sub>0 s" obtains X and t where "X \ Sols (\t. f) T S t\<^sub>0 s" and "X t = s'" and "t \ T" and "(\

X (down T t) \ {s. G s})" using assms unfolding g_orbital_def g_orbit_eq by auto no_notation g_orbit ("\") subsection \ Differential Invariants \ definition diff_invariant :: "('a \ bool) \ (('a::real_normed_vector) \ 'a) \ real set \ 'a set \ real \ ('a \ bool) \ bool" where "diff_invariant I f T S t\<^sub>0 G \ (\ \ (\

(g_orbital f G T S t\<^sub>0))) {s. I s} \ {s. I s}" lemma diff_invariant_eq: "diff_invariant I f T S t\<^sub>0 G = (\s. I s \ (\X\Sols (\t. f) T S t\<^sub>0 s. (\t\T.(\\\(down T t). G (X \)) \ I (X t))))" unfolding diff_invariant_def g_orbital_eq image_le_pred by auto lemma diff_inv_eq_inv_set: "diff_invariant I f T S t\<^sub>0 G = (\s. I s \ (g_orbital f G T S t\<^sub>0 s) \ {s. I s})" unfolding diff_invariant_eq g_orbital_eq image_le_pred by auto named_theorems diff_invariant_rules "rules for certifying differential invariants." lemma diff_invariant_eq_rule [diff_invariant_rules]: assumes Thyp: "is_interval T" "t\<^sub>0 \ T" and "\X. (D X = (\\. f (X \)) on T) \ (D (\\. \ (X \) - \ (X \)) = ((*\<^sub>R) 0) on T)" shows "diff_invariant (\s. \ s = \ s) f T S t\<^sub>0 G" proof(simp add: diff_invariant_eq ivp_sols_def, clarsimp) fix X \ assume tHyp:"\ \ T" and x_ivp:"D X = (\\. f (X \)) on T" "\ (X t\<^sub>0) = \ (X t\<^sub>0)" hence obs1: "\t\T. D (\\. \ (X \) - \ (X \)) \ (\\. \ *\<^sub>R 0) at t within T" using assms by (auto simp: has_vderiv_on_def has_vector_derivative_def) have obs2: "{t\<^sub>0--\} \ T" using closed_segment_subset_interval tHyp Thyp by blast hence "D (\\. \ (X \) - \ (X \)) = (\\. \ *\<^sub>R 0) on {t\<^sub>0--\}" using obs1 x_ivp by (auto intro!: has_derivative_subset[OF _ obs2] simp: has_vderiv_on_def has_vector_derivative_def) then obtain t where "t \ {t\<^sub>0--\}" and "\ (X \) - \ (X \) - (\ (X t\<^sub>0) - \ (X t\<^sub>0)) = (\ - t\<^sub>0) * t *\<^sub>R 0" using mvt_very_simple_closed_segmentE by blast thus "\ (X \) = \ (X \)" by (simp add: x_ivp(2)) qed lemma diff_invariant_leq_rule [diff_invariant_rules]: fixes \::"'a::banach \ real" assumes Thyp: "is_interval T" "t\<^sub>0 \ T" and "\X. (D X = (\\. f (X \)) on T) \ (\\\T. (\ > t\<^sub>0 \ \' (X \) \ \' (X \)) \ (\ < t\<^sub>0 \ \' (X \) \ \' (X \))) \ (D (\\. \ (X \) - \ (X \)) = (\\. \' (X \) - \' (X \)) on T)" shows "diff_invariant (\s. \ s \ \ s) f T S t\<^sub>0 G" proof(simp add: diff_invariant_eq ivp_sols_def, clarsimp) fix X \ assume "\ \ T" and x_ivp:"D X = (\\. f (X \)) on T" "\ (X t\<^sub>0) \ \ (X t\<^sub>0)" {assume "\ \ t\<^sub>0" hence primed: "\\. \ \ T \ \ > t\<^sub>0 \ \' (X \) \ \' (X \)" "\\. \ \ T \ \ < t\<^sub>0 \ \' (X \) \ \' (X \)" using x_ivp assms by auto have obs1: "\t\T. D (\\. \ (X \) - \ (X \)) \ (\\. \ *\<^sub>R (\' (X t) - \' (X t))) at t within T" using assms x_ivp by (auto simp: has_vderiv_on_def has_vector_derivative_def) have obs2: "{t\<^sub>0<--<\} \ T" "{t\<^sub>0--\} \ T" using \\ \ T\ Thyp \\ \ t\<^sub>0\ by (auto simp: convex_contains_open_segment is_interval_convex_1 closed_segment_subset_interval) hence "D (\\. \ (X \) - \ (X \)) = (\\. \' (X \) - \' (X \)) on {t\<^sub>0--\}" using obs1 x_ivp by (auto intro!: has_derivative_subset[OF _ obs2(2)] simp: has_vderiv_on_def has_vector_derivative_def) then obtain t where "t \ {t\<^sub>0<--<\}" and "(\ (X \) - \ (X \)) - (\ (X t\<^sub>0) - \ (X t\<^sub>0)) = (\\. \ * (\' (X t) - \' (X t))) (\ - t\<^sub>0)" using mvt_simple_closed_segmentE \\ \ t\<^sub>0\ by blast hence mvt: "\ (X \) - \ (X \) = (\ - t\<^sub>0) * (\' (X t) - \' (X t)) + (\ (X t\<^sub>0) - \ (X t\<^sub>0))" by force have "\ > t\<^sub>0 \ t > t\<^sub>0" "\ t\<^sub>0 \ \ \ t < t\<^sub>0" "t \ T" using \t \ {t\<^sub>0<--<\}\ obs2 unfolding open_segment_eq_real_ivl by auto moreover have "t > t\<^sub>0 \ (\' (X t) - \' (X t)) \ 0" "t < t\<^sub>0 \ (\' (X t) - \' (X t)) \ 0" using primed(1,2)[OF \t \ T\] by auto ultimately have "(\ - t\<^sub>0) * (\' (X t) - \' (X t)) \ 0" apply(case_tac "\ \ t\<^sub>0") by (force, auto simp: split_mult_pos_le) hence "(\ - t\<^sub>0) * (\' (X t) - \' (X t)) + (\ (X t\<^sub>0) - \ (X t\<^sub>0)) \ 0" using x_ivp(2) by auto hence "\ (X \) \ \ (X \)" using mvt by simp} thus "\ (X \) \ \ (X \)" using x_ivp by blast qed lemma diff_invariant_less_rule [diff_invariant_rules]: fixes \::"'a::banach \ real" assumes Thyp: "is_interval T" "t\<^sub>0 \ T" and "\ X. (D X = (\\. f (X \)) on T) \ (\\\T. (\ > t\<^sub>0 \ \' (X \) \ \' (X \)) \ (\ < t\<^sub>0 \ \' (X \) \ \' (X \))) \ (D (\\. \ (X \) - \ (X \)) = (\\. \' (X \) - \' (X \)) on T)" shows "diff_invariant (\s. \ s < \ s) f T S t\<^sub>0 G" proof(simp add: diff_invariant_eq ivp_sols_def, clarsimp) fix X \ assume "\ \ T" and x_ivp:"D X = (\\. f (X \)) on T" "\ (X t\<^sub>0) < \ (X t\<^sub>0)" {assume "\ \ t\<^sub>0" hence primed: "\\. \ \ T \ \ > t\<^sub>0 \ \' (X \) \ \' (X \)" "\\. \ \ T \ \ < t\<^sub>0 \ \' (X \) \ \' (X \)" using x_ivp assms by auto have obs1: "\t\T. D (\\. \ (X \) - \ (X \)) \ (\\. \ *\<^sub>R (\' (X t) - \' (X t))) at t within T" using assms x_ivp by (auto simp: has_vderiv_on_def has_vector_derivative_def) have obs2: "{t\<^sub>0<--<\} \ T" "{t\<^sub>0--\} \ T" using \\ \ T\ Thyp \\ \ t\<^sub>0\ by (auto simp: convex_contains_open_segment is_interval_convex_1 closed_segment_subset_interval) hence "D (\\. \ (X \) - \ (X \)) = (\\. \' (X \) - \' (X \)) on {t\<^sub>0--\}" using obs1 x_ivp by (auto intro!: has_derivative_subset[OF _ obs2(2)] simp: has_vderiv_on_def has_vector_derivative_def) then obtain t where "t \ {t\<^sub>0<--<\}" and "(\ (X \) - \ (X \)) - (\ (X t\<^sub>0) - \ (X t\<^sub>0)) = (\\. \ * (\' (X t) - \' (X t))) (\ - t\<^sub>0)" using mvt_simple_closed_segmentE \\ \ t\<^sub>0\ by blast hence mvt: "\ (X \) - \ (X \) = (\ - t\<^sub>0) * (\' (X t) - \' (X t)) + (\ (X t\<^sub>0) - \ (X t\<^sub>0))" by force have "\ > t\<^sub>0 \ t > t\<^sub>0" "\ t\<^sub>0 \ \ \ t < t\<^sub>0" "t \ T" using \t \ {t\<^sub>0<--<\}\ obs2 unfolding open_segment_eq_real_ivl by auto moreover have "t > t\<^sub>0 \ (\' (X t) - \' (X t)) \ 0" "t < t\<^sub>0 \ (\' (X t) - \' (X t)) \ 0" using primed(1,2)[OF \t \ T\] by auto ultimately have "(\ - t\<^sub>0) * (\' (X t) - \' (X t)) \ 0" apply(case_tac "\ \ t\<^sub>0") by (force, auto simp: split_mult_pos_le) hence "(\ - t\<^sub>0) * (\' (X t) - \' (X t)) + (\ (X t\<^sub>0) - \ (X t\<^sub>0)) > 0" using x_ivp(2) by auto hence "\ (X \) < \ (X \)" using mvt by simp} thus "\ (X \) < \ (X \)" using x_ivp by blast qed lemma diff_invariant_conj_rule [diff_invariant_rules]: assumes "diff_invariant I\<^sub>1 f T S t\<^sub>0 G" and "diff_invariant I\<^sub>2 f T S t\<^sub>0 G" shows "diff_invariant (\s. I\<^sub>1 s \ I\<^sub>2 s) f T S t\<^sub>0 G" using assms unfolding diff_invariant_def by auto lemma diff_invariant_disj_rule [diff_invariant_rules]: assumes "diff_invariant I\<^sub>1 f T S t\<^sub>0 G" and "diff_invariant I\<^sub>2 f T S t\<^sub>0 G" shows "diff_invariant (\s. I\<^sub>1 s \ I\<^sub>2 s) f T S t\<^sub>0 G" using assms unfolding diff_invariant_def by auto subsection \ Picard-Lindeloef \ text\ A locale with the assumptions of Picard-Lindeloef theorem. It extends @{term "ll_on_open_it"} by providing an initial time @{term "t\<^sub>0 \ T"}.\ locale picard_lindeloef = fixes f::"real \ ('a::{heine_borel,banach}) \ 'a" and T::"real set" and S::"'a set" and t\<^sub>0::real assumes open_domain: "open T" "open S" and interval_time: "is_interval T" and init_time: "t\<^sub>0 \ T" and cont_vec_field: "\s \ S. continuous_on T (\t. f t s)" and lipschitz_vec_field: "local_lipschitz T S f" begin sublocale ll_on_open_it T f S t\<^sub>0 by (unfold_locales) (auto simp: cont_vec_field lipschitz_vec_field interval_time open_domain) lemmas subintervalI = closed_segment_subset_domain lemma csols_eq: "csols t\<^sub>0 s = {(X, t). t \ T \ X \ Sols f {t\<^sub>0--t} S t\<^sub>0 s}" unfolding ivp_sols_def csols_def solves_ode_def using subintervalI[OF init_time] by auto abbreviation "ex_ivl s \ existence_ivl t\<^sub>0 s" lemma unique_solution: assumes xivp: "D X = (\t. f t (X t)) on {t\<^sub>0--t}" "X t\<^sub>0 = s" "X \ {t\<^sub>0--t} \ S" and "t \ T" and yivp: "D Y = (\t. f t (Y t)) on {t\<^sub>0--t}" "Y t\<^sub>0 = s" "Y \ {t\<^sub>0--t} \ S" and "s \ S" shows "X t = Y t" proof- have "(X, t) \ csols t\<^sub>0 s" using xivp \t \ T\ unfolding csols_eq ivp_sols_def by auto hence ivl_fact: "{t\<^sub>0--t} \ ex_ivl s" unfolding existence_ivl_def by auto have obs: "\z T'. t\<^sub>0 \ T' \ is_interval T' \ T' \ ex_ivl s \ (z solves_ode f) T' S \ z t\<^sub>0 = flow t\<^sub>0 s t\<^sub>0 \ (\t\T'. z t = flow t\<^sub>0 s t)" using flow_usolves_ode[OF init_time \s \ S\] unfolding usolves_ode_from_def by blast have "\\\{t\<^sub>0--t}. X \ = flow t\<^sub>0 s \" using obs[of "{t\<^sub>0--t}" X] xivp ivl_fact flow_initial_time[OF init_time \s \ S\] unfolding solves_ode_def by simp also have "\\\{t\<^sub>0--t}. Y \ = flow t\<^sub>0 s \" using obs[of "{t\<^sub>0--t}" Y] yivp ivl_fact flow_initial_time[OF init_time \s \ S\] unfolding solves_ode_def by simp ultimately show "X t = Y t" by auto qed lemma solution_eq_flow: assumes xivp: "D X = (\t. f t (X t)) on ex_ivl s" "X t\<^sub>0 = s" "X \ ex_ivl s \ S" and "t \ ex_ivl s" and "s \ S" shows "X t = flow t\<^sub>0 s t" proof- have obs: "\z T'. t\<^sub>0 \ T' \ is_interval T' \ T' \ ex_ivl s \ (z solves_ode f) T' S \ z t\<^sub>0 = flow t\<^sub>0 s t\<^sub>0 \ (\t\T'. z t = flow t\<^sub>0 s t)" using flow_usolves_ode[OF init_time \s \ S\] unfolding usolves_ode_from_def by blast have "\\\ex_ivl s. X \ = flow t\<^sub>0 s \" using obs[of "ex_ivl s" X] existence_ivl_initial_time[OF init_time \s \ S\] xivp flow_initial_time[OF init_time \s \ S\] unfolding solves_ode_def by simp thus "X t = flow t\<^sub>0 s t" by (auto simp: \t \ ex_ivl s\) qed end lemma local_lipschitz_add: fixes f1 f2 :: "real \ 'a::banach \ 'a" assumes "local_lipschitz T S f1" and "local_lipschitz T S f2" shows "local_lipschitz T S (\t s. f1 t s + f2 t s)" proof(unfold local_lipschitz_def, clarsimp) fix s and t assume "s \ S" and "t \ T" obtain \\<^sub>1 L1 where "\\<^sub>1 > 0" and L1: "\\. \\cball t \\<^sub>1 \ T \ L1-lipschitz_on (cball s \\<^sub>1 \ S) (f1 \)" using local_lipschitzE[OF assms(1) \t \ T\ \s \ S\] by blast obtain \\<^sub>2 L2 where "\\<^sub>2 > 0" and L2: "\\. \\cball t \\<^sub>2 \ T \ L2-lipschitz_on (cball s \\<^sub>2 \ S) (f2 \)" using local_lipschitzE[OF assms(2) \t \ T\ \s \ S\] by blast have ballH: "cball s (min \\<^sub>1 \\<^sub>2) \ S \ cball s \\<^sub>1 \ S" "cball s (min \\<^sub>1 \\<^sub>2) \ S \ cball s \\<^sub>2 \ S" by auto have obs1: "\\\cball t \\<^sub>1 \ T. L1-lipschitz_on (cball s (min \\<^sub>1 \\<^sub>2) \ S) (f1 \)" using lipschitz_on_subset[OF L1 ballH(1)] by blast also have obs2: "\\\cball t \\<^sub>2 \ T. L2-lipschitz_on (cball s (min \\<^sub>1 \\<^sub>2) \ S) (f2 \)" using lipschitz_on_subset[OF L2 ballH(2)] by blast ultimately have "\\\cball t (min \\<^sub>1 \\<^sub>2) \ T. (L1 + L2)-lipschitz_on (cball s (min \\<^sub>1 \\<^sub>2) \ S) (\s. f1 \ s + f2 \ s)" using lipschitz_on_add by fastforce thus "\u>0. \L. \t\cball t u \ T. L-lipschitz_on (cball s u \ S) (\s. f1 t s + f2 t s)" apply(rule_tac x="min \\<^sub>1 \\<^sub>2" in exI) using \\\<^sub>1 > 0\ \\\<^sub>2 > 0\ by force qed lemma picard_lindeloef_add: "picard_lindeloef f1 T S t\<^sub>0 \ picard_lindeloef f2 T S t\<^sub>0 \ picard_lindeloef (\t s. f1 t s + f2 t s) T S t\<^sub>0" unfolding picard_lindeloef_def apply(clarsimp, rule conjI) using continuous_on_add apply fastforce using local_lipschitz_add by blast lemma picard_lindeloef_constant: "picard_lindeloef (\t s. c) UNIV UNIV t\<^sub>0" apply(unfold_locales, simp_all add: local_lipschitz_def lipschitz_on_def, clarsimp) by (rule_tac x=1 in exI, clarsimp, rule_tac x="1/2" in exI, simp) subsection \ Flows for ODEs \ text\ A locale designed for verification of hybrid systems. The user can select the interval of existence and the defining flow equation via the variables @{term "T"} and @{term "\"}.\ locale local_flow = picard_lindeloef "(\ t. f)" T S 0 for f::"'a::{heine_borel,banach} \ 'a" and T S L + fixes \ :: "real \ 'a \ 'a" assumes ivp: "\ t s. t \ T \ s \ S \ D (\t. \ t s) = (\t. f (\ t s)) on {0--t}" "\ s. s \ S \ \ 0 s = s" "\ t s. t \ T \ s \ S \ (\t. \ t s) \ {0--t} \ S" begin lemma in_ivp_sols_ivl: assumes "t \ T" "s \ S" shows "(\t. \ t s) \ Sols (\t. f) {0--t} S 0 s" apply(rule ivp_solsI) using ivp assms by auto lemma eq_solution_ivl: assumes xivp: "D X = (\t. f (X t)) on {0--t}" "X 0 = s" "X \ {0--t} \ S" and indom: "t \ T" "s \ S" shows "X t = \ t s" apply(rule unique_solution[OF xivp \t \ T\]) using \s \ S\ ivp indom by auto lemma ex_ivl_eq: assumes "s \ S" shows "ex_ivl s = T" using existence_ivl_subset[of s] apply safe unfolding existence_ivl_def csols_eq using in_ivp_sols_ivl[OF _ assms] by blast lemma has_derivative_on_open1: assumes "t > 0" "t \ T" "s \ S" obtains B where "t \ B" and "open B" and "B \ T" and "D (\\. \ \ s) \ (\\. \ *\<^sub>R f (\ t s)) at t within B" proof- obtain r::real where rHyp: "r > 0" "ball t r \ T" using open_contains_ball_eq open_domain(1) \t \ T\ by blast moreover have "t + r/2 > 0" using \r > 0\ \t > 0\ by auto moreover have "{0--t} \ T" using subintervalI[OF init_time \t \ T\] . ultimately have subs: "{0<-- T" unfolding abs_le_eq abs_le_eq real_ivl_eqs[OF \t > 0\] real_ivl_eqs[OF \t + r/2 > 0\] by clarify (case_tac "t < x", simp_all add: cball_def ball_def dist_norm subset_eq field_simps) have "t + r/2 \ T" using rHyp unfolding real_ivl_eqs[OF rHyp(1)] by (simp add: subset_eq) hence "{0--t + r/2} \ T" using subintervalI[OF init_time] by blast hence "(D (\t. \ t s) = (\t. f (\ t s)) on {0--(t + r/2)})" using ivp(1)[OF _ \s \ S\] by auto hence vderiv: "(D (\t. \ t s) = (\t. f (\ t s)) on {0<--t + r/2 > 0\] by auto have "t \ {0<--t + r/2 > 0\] using rHyp \t > 0\ by simp moreover have "D (\\. \ \ s) \ (\\. \ *\<^sub>R f (\ t s)) (at t within {0<--t + r/2 > 0\] by simp ultimately show ?thesis using subs that by blast qed lemma has_derivative_on_open2: assumes "t < 0" "t \ T" "s \ S" obtains B where "t \ B" and "open B" and "B \ T" and "D (\\. \ \ s) \ (\\. \ *\<^sub>R f (\ t s)) at t within B" proof- obtain r::real where rHyp: "r > 0" "ball t r \ T" using open_contains_ball_eq open_domain(1) \t \ T\ by blast moreover have "t - r/2 < 0" using \r > 0\ \t < 0\ by auto moreover have "{0--t} \ T" using subintervalI[OF init_time \t \ T\] . ultimately have subs: "{0<-- T" unfolding open_segment_eq_real_ivl closed_segment_eq_real_ivl real_ivl_eqs[OF rHyp(1)] by(auto simp: subset_eq) have "t - r/2 \ T" using rHyp unfolding real_ivl_eqs by (simp add: subset_eq) hence "{0--t - r/2} \ T" using subintervalI[OF init_time] by blast hence "(D (\t. \ t s) = (\t. f (\ t s)) on {0--(t - r/2)})" using ivp(1)[OF _ \s \ S\] by auto hence vderiv: "(D (\t. \ t s) = (\t. f (\ t s)) on {0<-- {0<--t < 0\ by simp moreover have "D (\\. \ \ s) \ (\\. \ *\<^sub>R f (\ t s)) (at t within {0<-- S" obtains B where "0 \ B" and "open B" and "B \ T" and "D (\\. \ \ s) \ (\\. \ *\<^sub>R f (\ 0 s)) at 0 within B" proof- obtain r::real where rHyp: "r > 0" "ball 0 r \ T" using open_contains_ball_eq open_domain(1) init_time by blast hence "r/2 \ T" "-r/2 \ T" "r/2 > 0" unfolding real_ivl_eqs by auto hence subs: "{0--r/2} \ T" "{0--(-r/2)} \ T" using subintervalI[OF init_time] by auto hence "(D (\t. \ t s) = (\t. f (\ t s)) on {0--r/2})" "(D (\t. \ t s) = (\t. f (\ t s)) on {0--(-r/2)})" using ivp(1)[OF _ \s \ S\] by auto also have "{0--r/2} = {0--r/2} \ closure {0--r/2} \ closure {0--(-r/2)}" "{0--(-r/2)} = {0--(-r/2)} \ closure {0--r/2} \ closure {0--(-r/2)}" unfolding closed_segment_eq_real_ivl \r/2 > 0\ by auto ultimately have vderivs: "(D (\t. \ t s) = (\t. f (\ t s)) on {0--r/2} \ closure {0--r/2} \ closure {0--(-r/2)})" "(D (\t. \ t s) = (\t. f (\ t s)) on {0--(-r/2)} \ closure {0--r/2} \ closure {0--(-r/2)})" unfolding closed_segment_eq_real_ivl \r/2 > 0\ by auto have obs: "0 \ {-r/2<--r/2 > 0\ by auto have union: "{-r/2--r/2} = {0--r/2} \ {0--(-r/2)}" unfolding closed_segment_eq_real_ivl by auto hence "(D (\t. \ t s) = (\t. f (\ t s)) on {-r/2--r/2})" using has_vderiv_on_union[OF vderivs] by simp hence "(D (\t. \ t s) = (\t. f (\ t s)) on {-r/2<--\. \ \ s) \ (\\. \ *\<^sub>R f (\ 0 s)) (at 0 within {-r/2<-- T" using subs union segment_open_subset_closed by blast ultimately show ?thesis using obs that by blast qed lemma has_derivative_on_open: assumes "t \ T" "s \ S" obtains B where "t \ B" and "open B" and "B \ T" and "D (\\. \ \ s) \ (\\. \ *\<^sub>R f (\ t s)) at t within B" apply(subgoal_tac "t < 0 \ t = 0 \ t > 0") using has_derivative_on_open1[OF _ assms] has_derivative_on_open2[OF _ assms] has_derivative_on_open3[OF \s \ S\] by blast force lemma in_domain: assumes "s \ S" shows "(\t. \ t s) \ T \ S" unfolding ex_ivl_eq[symmetric] existence_ivl_def using local.mem_existence_ivl_subset ivp(3)[OF _ assms] by blast lemma has_vderiv_on_domain: assumes "s \ S" shows "D (\t. \ t s) = (\t. f (\ t s)) on T" proof(unfold has_vderiv_on_def has_vector_derivative_def, clarsimp) fix t assume "t \ T" then obtain B where "t \ B" and "open B" and "B \ T" and Dhyp: "D (\t. \ t s) \ (\\. \ *\<^sub>R f (\ t s)) at t within B" using assms has_derivative_on_open[OF \t \ T\] by blast hence "t \ interior B" using interior_eq by auto thus "D (\t. \ t s) \ (\\. \ *\<^sub>R f (\ t s)) at t within T" using has_derivative_at_within_mono[OF _ \B \ T\ Dhyp] by blast qed lemma in_ivp_sols: assumes "s \ S" shows "(\t. \ t s) \ Sols (\t. f) T S 0 s" using has_vderiv_on_domain ivp(2) in_domain apply(rule ivp_solsI) using assms by auto lemma eq_solution: assumes "X \ Sols (\t. f) T S 0 s" and "t \ T" and "s \ S" shows "X t = \ t s" proof- have "D X = (\t. f (X t)) on (ex_ivl s)" and "X 0 = s" and "X \ (ex_ivl s) \ S" using ivp_solsD[OF assms(1)] unfolding ex_ivl_eq[OF \s \ S\] by auto note solution_eq_flow[OF this] hence "X t = flow 0 s t" unfolding ex_ivl_eq[OF \s \ S\] using assms by blast also have "\ t s = flow 0 s t" apply(rule solution_eq_flow ivp) apply(simp_all add: assms(2,3) ivp(2)[OF \s \ S\]) unfolding ex_ivl_eq[OF \s \ S\] by (auto simp: has_vderiv_on_domain assms in_domain) ultimately show "X t = \ t s" by simp qed lemma ivp_sols_collapse: assumes "T = UNIV" and "s \ S" shows "Sols (\t. f) T S 0 s = {(\t. \ t s)}" using in_ivp_sols eq_solution assms by auto lemma additive_in_ivp_sols: assumes "s \ S" and "\

(\\. \ + t) T \ T" shows "(\\. \ (\ + t) s) \ Sols (\t. f) T S 0 (\ (0 + t) s)" apply(rule ivp_solsI, rule vderiv_on_compose_add) using has_vderiv_on_domain has_vderiv_on_subset assms apply blast using in_domain assms by auto lemma is_monoid_action: assumes "s \ S" and "T = UNIV" shows "\ 0 s = s" and "\ (t\<^sub>1 + t\<^sub>2) s = \ t\<^sub>1 (\ t\<^sub>2 s)" proof- show "\ 0 s = s" using ivp assms by simp have "\ (0 + t\<^sub>2) s = \ t\<^sub>2 s" by simp also have "\ t\<^sub>2 s \ S" using in_domain assms by auto finally show "\ (t\<^sub>1 + t\<^sub>2) s = \ t\<^sub>1 (\ t\<^sub>2 s)" using eq_solution[OF additive_in_ivp_sols] assms by auto qed definition orbit :: "'a \ 'a set" ("\\<^sup>\") where "\\<^sup>\ s = g_orbital f (\s. True) T S 0 s" lemma orbit_eq[simp]: assumes "s \ S" shows "\\<^sup>\ s = {\ t s| t. t \ T}" using eq_solution assms unfolding orbit_def g_orbital_eq ivp_sols_def by(auto intro!: has_vderiv_on_domain ivp(2) in_domain) lemma g_orbital_collapses: assumes "s \ S" shows "g_orbital f G T S 0 s = {\ t s| t. t \ T \ (\\\down T t. G (\ \ s))}" proof(rule subset_antisym, simp_all only: subset_eq) let ?gorbit = "{\ t s |t. t \ T \ (\\\down T t. G (\ \ s))}" {fix s' assume "s' \ g_orbital f G T S 0 s" then obtain X and t where x_ivp:"X \ Sols (\t. f) T S 0 s" and "X t = s'" and "t \ T" and guard:"(\

X (down T t) \ {s. G s})" unfolding g_orbital_def g_orbit_eq by auto have obs:"\\\(down T t). X \ = \ \ s" using eq_solution[OF x_ivp _ assms] by blast hence "\

(\t. \ t s) (down T t) \ {s. G s}" using guard by auto also have "\ t s = X t" using eq_solution[OF x_ivp \t \ T\ assms] by simp ultimately have "s' \ ?gorbit" using \X t = s'\ \t \ T\ by auto} thus "\s'\ g_orbital f G T S 0 s. s' \ ?gorbit" by blast next let ?gorbit = "{\ t s |t. t \ T \ (\\\down T t. G (\ \ s))}" {fix s' assume "s' \ ?gorbit" then obtain t where "\

(\t. \ t s) (down T t) \ {s. G s}" and "t \ T" and "\ t s = s'" by blast hence "s' \ g_orbital f G T S 0 s" using assms by(auto intro!: g_orbitalI in_ivp_sols)} thus "\s'\?gorbit. s' \ g_orbital f G T S 0 s" by blast qed end lemma line_is_local_flow: "0 \ T \ is_interval T \ open T \ local_flow (\ s. c) T UNIV (\ t s. s + t *\<^sub>R c)" apply(unfold_locales, simp_all add: local_lipschitz_def lipschitz_on_def, clarsimp) apply(rule_tac x=1 in exI, clarsimp, rule_tac x="1/2" in exI, simp) - apply(rule_tac f'1="\ s. 0" and g'1="\ s. c" in derivative_intros(191)) + apply(rule_tac f'1="\ s. 0" and g'1="\ s. c" in has_vderiv_on_add[THEN has_vderiv_on_eq_rhs]) apply(rule derivative_intros, simp)+ by simp_all end \ No newline at end of file diff --git a/thys/Hybrid_Systems_VCs/HS_Preliminaries.thy b/thys/Hybrid_Systems_VCs/HS_Preliminaries.thy --- a/thys/Hybrid_Systems_VCs/HS_Preliminaries.thy +++ b/thys/Hybrid_Systems_VCs/HS_Preliminaries.thy @@ -1,315 +1,317 @@ (* Title: Preliminaries for hybrid systems verification Author: Jonathan Julián Huerta y Munive, 2019 Maintainer: Jonathan Julián Huerta y Munive *) section \ Hybrid Systems Preliminaries \ text \Hybrid systems combine continuous dynamics with discrete control. This section contains auxiliary lemmas for verification of hybrid systems.\ theory HS_Preliminaries imports "Ordinary_Differential_Equations.Picard_Lindeloef_Qualitative" begin subsection \ Real numbers \ lemma abs_le_eq: shows "(r::real) > 0 \ (\x\ < r) = (-r < x \ x < r)" and "(r::real) > 0 \ (\x\ \ r) = (-r \ x \ x \ r)" by linarith+ lemma real_ivl_eqs: assumes "0 < r" shows "ball x r = {x-r<--< x+r}" and "{x-r<--< x+r} = {x-r<..< x+r}" and "ball (r / 2) (r / 2) = {0<--2 + (x * sin t + y * cos t)\<^sup>2 = x\<^sup>2 + y\<^sup>2" and "(x * cos t + y * sin t)\<^sup>2 + (y * cos t - x * sin t)\<^sup>2 = x\<^sup>2 + y\<^sup>2" proof- have "(x * cos t - y * sin t)\<^sup>2 = x\<^sup>2 * (cos t)\<^sup>2 + y\<^sup>2 * (sin t)\<^sup>2 - 2 * (x * cos t) * (y * sin t)" by(simp add: power2_diff power_mult_distrib) also have "(x * sin t + y * cos t)\<^sup>2 = y\<^sup>2 * (cos t)\<^sup>2 + x\<^sup>2 * (sin t)\<^sup>2 + 2 * (x * cos t) * (y * sin t)" by(simp add: power2_sum power_mult_distrib) ultimately show "(x * cos t - y * sin t)\<^sup>2 + (x * sin t + y * cos t)\<^sup>2 = x\<^sup>2 + y\<^sup>2" by (simp add: Groups.mult_ac(2) Groups.mult_ac(3) right_diff_distrib sin_squared_eq) thus "(x * cos t + y * sin t)\<^sup>2 + (y * cos t - x * sin t)\<^sup>2 = x\<^sup>2 + y\<^sup>2" by (simp add: add.commute add.left_commute power2_diff power2_sum) qed subsection \ Single variable derivatives \ notation has_derivative ("(1(D _ \ (_))/ _)" [65,65] 61) notation has_vderiv_on ("(1 D _ = (_)/ on _)" [65,65] 61) notation norm ("(1\_\)" [65] 61) +thy_deps + named_theorems poly_derivatives "compilation of optimised miscellaneous derivative rules." declare has_vderiv_on_const [poly_derivatives] and has_vderiv_on_id [poly_derivatives] - and derivative_intros(191) [poly_derivatives] - and derivative_intros(192) [poly_derivatives] - and derivative_intros(194) [poly_derivatives] + and has_vderiv_on_add[THEN has_vderiv_on_eq_rhs, poly_derivatives] + and has_vderiv_on_diff[THEN has_vderiv_on_eq_rhs, poly_derivatives] + and has_vderiv_on_mult[THEN has_vderiv_on_eq_rhs, poly_derivatives] lemma has_vderiv_on_compose_eq: assumes "D f = f' on g ` T" and " D g = g' on T" and "h = (\x. g' x *\<^sub>R f' (g x))" shows "D (\t. f (g t)) = h on T" apply(subst ssubst[of h], simp) using assms has_vderiv_on_compose by auto lemma vderiv_on_compose_add [derivative_intros]: assumes "D x = x' on (\\. \ + t) ` T" shows "D (\\. x (\ + t)) = (\\. x' (\ + t)) on T" apply(rule has_vderiv_on_compose_eq[OF assms]) by(auto intro: derivative_intros) lemma has_vderiv_on_divide_cnst: "a \ 0 \ D (\t. t/a) = (\t. 1/a) on T" unfolding has_vderiv_on_def has_vector_derivative_def apply clarify - apply(rule_tac f'1="\t. t" and g'1="\ x. 0" in derivative_eq_intros(18)) + apply(rule_tac f'1="\t. t" and g'1="\ x. 0" in has_derivative_divide[THEN has_derivative_eq_rhs]) by(auto intro: derivative_eq_intros) lemma has_vderiv_on_power: "n \ 1 \ D (\t. t ^ n) = (\t. n * (t ^ (n - 1))) on T" unfolding has_vderiv_on_def has_vector_derivative_def by (auto intro!: derivative_eq_intros) lemma has_vderiv_on_exp: "D (\t. exp t) = (\t. exp t) on T" unfolding has_vderiv_on_def has_vector_derivative_def by (auto intro: derivative_intros) lemma has_vderiv_on_cos_comp: "D (f::real \ real) = f' on T \ D (\t. cos (f t)) = (\t. - (f' t) * sin (f t)) on T" apply(rule has_vderiv_on_compose_eq[of "\t. cos t"]) unfolding has_vderiv_on_def has_vector_derivative_def apply clarify by(auto intro!: derivative_eq_intros simp: fun_eq_iff) lemma has_vderiv_on_sin_comp: "D (f::real \ real) = f' on T \ D (\t. sin (f t)) = (\t. (f' t) * cos (f t)) on T" apply(rule has_vderiv_on_compose_eq[of "\t. sin t"]) unfolding has_vderiv_on_def has_vector_derivative_def apply clarify by(auto intro!: derivative_eq_intros simp: fun_eq_iff) lemma has_vderiv_on_exp_comp: "D (f::real \ real) = f' on T \ D (\t. exp (f t)) = (\t. (f' t) * exp (f t)) on T" apply(rule has_vderiv_on_compose_eq[of "\t. exp t"]) by (rule has_vderiv_on_exp, simp_all add: mult.commute) lemma vderiv_uminus_intro [poly_derivatives]: "D f = f' on T \ g = (\t. - f' t) \ D (\t. - f t) = g on T" using has_vderiv_on_uminus by auto lemma vderiv_div_cnst_intro [poly_derivatives]: assumes "(a::real) \ 0" and "D f = f' on T" and "g = (\t. (f' t)/a)" shows "D (\t. (f t)/a) = g on T" apply(rule has_vderiv_on_compose_eq[of "\t. t/a" "\t. 1/a"]) using assms by(auto intro: has_vderiv_on_divide_cnst) lemma vderiv_npow_intro [poly_derivatives]: fixes f::"real \ real" assumes "n \ 1" and "D f = f' on T" and "g = (\t. n * (f' t) * (f t)^(n-1))" shows "D (\t. (f t)^n) = g on T" apply(rule has_vderiv_on_compose_eq[of "\t. t^n"]) using assms(1) apply(rule has_vderiv_on_power) using assms by auto lemma vderiv_cos_intro [poly_derivatives]: assumes "D (f::real \ real) = f' on T" and "g = (\t. - (f' t) * sin (f t))" shows "D (\t. cos (f t)) = g on T" using assms and has_vderiv_on_cos_comp by auto lemma vderiv_sin_intro [poly_derivatives]: assumes "D (f::real \ real) = f' on T" and "g = (\t. (f' t) * cos (f t))" shows "D (\t. sin (f t)) = g on T" using assms and has_vderiv_on_sin_comp by auto lemma vderiv_exp_intro [poly_derivatives]: assumes "D (f::real \ real) = f' on T" and "g = (\t. (f' t) * exp (f t))" shows "D (\t. exp (f t)) = g on T" using assms and has_vderiv_on_exp_comp by auto \ \Examples for checking derivatives\ lemma "D (\t. a * t\<^sup>2 / 2 + v * t + x) = (\t. a * t + v) on T" by(auto intro!: poly_derivatives) lemma "D (\t. v * t - a * t\<^sup>2 / 2 + x) = (\x. v - a * x) on T" by(auto intro!: poly_derivatives) lemma "c \ 0 \ D (\t. a5 * t^5 + a3 * (t^3 / c) - a2 * exp (t^2) + a1 * cos t + a0) = (\t. 5 * a5 * t^4 + 3 * a3 * (t^2 / c) - 2 * a2 * t * exp (t^2) - a1 * sin t) on T" by(auto intro!: poly_derivatives) lemma "c \ 0 \ D (\t. - a3 * exp (t^3 / c) + a1 * sin t + a2 * t^2) = (\t. a1 * cos t + 2 * a2 * t - 3 * a3 * t^2 / c * exp (t^3 / c)) on T" apply(intro poly_derivatives) using poly_derivatives(1,2) by force+ lemma "c \ 0 \ D (\t. exp (a * sin (cos (t^4) / c))) = (\t. - 4 * a * t^3 * sin (t^4) / c * cos (cos (t^4) / c) * exp (a * sin (cos (t^4) / c))) on T" apply(intro poly_derivatives) using poly_derivatives(1,2) by force+ subsection \ Filters \ lemma eventually_at_within_mono: assumes "t \ interior T" and "T \ S" and "eventually P (at t within T)" shows "eventually P (at t within S)" by (meson assms eventually_within_interior interior_mono subsetD) lemma netlimit_at_within_mono: fixes t::"'a::{perfect_space,t2_space}" assumes "t \ interior T" and "T \ S" shows "netlimit (at t within S) = t" using assms(1) interior_mono[OF \T \ S\] netlimit_within_interior by auto lemma has_derivative_at_within_mono: assumes "(t::real) \ interior T" and "T \ S" and "D f \ f' at t within T" shows "D f \ f' at t within S" using assms(3) apply(unfold has_derivative_def tendsto_iff, safe) unfolding netlimit_at_within_mono[OF assms(1,2)] netlimit_within_interior[OF assms(1)] by (rule eventually_at_within_mono[OF assms(1,2)]) simp lemma eventually_all_finite2: fixes P :: "('a::finite) \ 'b \ bool" assumes h:"\i. eventually (P i) F" shows "eventually (\x. \i. P i x) F" proof(unfold eventually_def) let ?F = "Rep_filter F" have obs: "\i. ?F (P i)" using h by auto have "?F (\x. \i \ UNIV. P i x)" apply(rule finite_induct) by(auto intro: eventually_conj simp: obs h) thus "?F (\x. \i. P i x)" by simp qed lemma eventually_all_finite_mono: fixes P :: "('a::finite) \ 'b \ bool" assumes h1: "\i. eventually (P i) F" and h2: "\x. (\i. (P i x)) \ Q x" shows "eventually Q F" proof- have "eventually (\x. \i. P i x) F" using h1 eventually_all_finite2 by blast thus "eventually Q F" unfolding eventually_def using h2 eventually_mono by auto qed subsection \ Multivariable derivatives \ lemma frechet_vec_lambda: fixes f::"real \ ('a::banach)^('m::finite)" and x::real and T::"real set" defines "x\<^sub>0 \ netlimit (at x within T)" and "m \ real CARD('m)" assumes "\i. ((\y. (f y $ i - f x\<^sub>0 $ i - (y - x\<^sub>0) *\<^sub>R f' x $ i) /\<^sub>R (\y - x\<^sub>0\)) \ 0) (at x within T)" shows "((\y. (f y - f x\<^sub>0 - (y - x\<^sub>0) *\<^sub>R f' x) /\<^sub>R (\y - x\<^sub>0\)) \ 0) (at x within T)" proof(simp add: tendsto_iff, clarify) fix \::real assume "0 < \" let "?\" = "\y. y - x\<^sub>0" and "?\f" = "\y. f y - f x\<^sub>0" let "?P" = "\i e y. inverse \?\ y\ * (\f y $ i - f x\<^sub>0 $ i - ?\ y *\<^sub>R f' x $ i\) < e" and "?Q" = "\y. inverse \?\ y\ * (\?\f y - ?\ y *\<^sub>R f' x\) < \" have "0 < \ / sqrt m" using \0 < \\ by (auto simp: assms) hence "\i. eventually (\y. ?P i (\ / sqrt m) y) (at x within T)" using assms unfolding tendsto_iff by simp thus "eventually ?Q (at x within T)" proof(rule eventually_all_finite_mono, simp add: norm_vec_def L2_set_def, clarify) fix t::real let ?c = "inverse \t - x\<^sub>0\" and "?u t" = "\i. f t $ i - f x\<^sub>0 $ i - ?\ t *\<^sub>R f' x $ i" assume hyp:"\i. ?c * (\?u t i\) < \ / sqrt m" hence "\i. (?c *\<^sub>R (\?u t i\))\<^sup>2 < (\ / sqrt m)\<^sup>2" by (simp add: power_strict_mono) hence "\i. ?c\<^sup>2 * ((\?u t i\))\<^sup>2 < \\<^sup>2 / m" by (simp add: power_mult_distrib power_divide assms) hence "\i. ?c\<^sup>2 * ((\?u t i\))\<^sup>2 < \\<^sup>2 / m" by (auto simp: assms) also have "({}::'m set) \ UNIV \ finite (UNIV :: 'm set)" by simp ultimately have "(\i\UNIV. ?c\<^sup>2 * ((\?u t i\))\<^sup>2) < (\(i::'m)\UNIV. \\<^sup>2 / m)" by (metis (lifting) sum_strict_mono) moreover have "?c\<^sup>2 * (\i\UNIV. (\?u t i\)\<^sup>2) = (\i\UNIV. ?c\<^sup>2 * (\?u t i\)\<^sup>2)" using sum_distrib_left by blast ultimately have "?c\<^sup>2 * (\i\UNIV. (\?u t i\)\<^sup>2) < \\<^sup>2" by (simp add: assms) hence "sqrt (?c\<^sup>2 * (\i\UNIV. (\?u t i\)\<^sup>2)) < sqrt (\\<^sup>2)" using real_sqrt_less_iff by blast also have "... = \" using \0 < \\ by auto moreover have "?c * sqrt (\i\UNIV. (\?u t i\)\<^sup>2) = sqrt (?c\<^sup>2 * (\i\UNIV. (\?u t i\)\<^sup>2))" by (simp add: real_sqrt_mult) ultimately show "?c * sqrt (\i\UNIV. (\?u t i\)\<^sup>2) < \" by simp qed qed lemma frechet_vec_nth: fixes f::"real \ ('a::real_normed_vector)^'m" and x::real and T::"real set" defines "x\<^sub>0 \ netlimit (at x within T)" assumes "((\y. (f y - f x\<^sub>0 - (y - x\<^sub>0) *\<^sub>R f' x) /\<^sub>R (\y - x\<^sub>0\)) \ 0) (at x within T)" shows "((\y. (f y $ i - f x\<^sub>0 $ i - (y - x\<^sub>0) *\<^sub>R f' x $ i) /\<^sub>R (\y - x\<^sub>0\)) \ 0) (at x within T)" proof(unfold tendsto_iff dist_norm, clarify) let "?\" = "\y. y - x\<^sub>0" and "?\f" = "\y. f y - f x\<^sub>0" fix \::real assume "0 < \" let "?P" = "\y. \(?\f y - ?\ y *\<^sub>R f' x) /\<^sub>R (\?\ y\) - 0\ < \" and "?Q" = "\y. \(f y $ i - f x\<^sub>0 $ i - ?\ y *\<^sub>R f' x $ i) /\<^sub>R (\?\ y\) - 0\ < \" have "eventually ?P (at x within T)" using \0 < \\ assms unfolding tendsto_iff by auto thus "eventually ?Q (at x within T)" proof(rule_tac P="?P" in eventually_mono, simp_all) fix y let "?u i" = "f y $ i - f x\<^sub>0 $ i - ?\ y *\<^sub>R f' x $ i" assume hyp:"inverse \?\ y\ * (\?\f y - ?\ y *\<^sub>R f' x\) < \" have "\(?\f y - ?\ y *\<^sub>R f' x) $ i\ \ \?\f y - ?\ y *\<^sub>R f' x\" using Finite_Cartesian_Product.norm_nth_le by blast also have "\?u i\ = \(?\f y - ?\ y *\<^sub>R f' x) $ i\" by simp ultimately have "\?u i\ \ \?\f y - ?\ y *\<^sub>R f' x\" by linarith hence "inverse \?\ y\ * (\?u i\) \ inverse \?\ y\ * (\?\f y - ?\ y *\<^sub>R f' x\)" by (simp add: mult_left_mono) thus "inverse \?\ y\ * (\f y $ i - f x\<^sub>0 $ i - ?\ y *\<^sub>R f' x $ i\) < \" using hyp by linarith qed qed lemma has_derivative_vec_lambda: fixes f::"real \ ('a::banach)^('n::finite)" assumes "\i. D (\t. f t $ i) \ (\ h. h *\<^sub>R f' x $ i) (at x within T)" shows "D f \ (\h. h *\<^sub>R f' x) at x within T" apply(unfold has_derivative_def, safe) apply(force simp: bounded_linear_def bounded_linear_axioms_def) using assms frechet_vec_lambda[of x T ] unfolding has_derivative_def by auto lemma has_derivative_vec_nth: assumes "D f \ (\h. h *\<^sub>R f' x) at x within T" shows "D (\t. f t $ i) \ (\h. h *\<^sub>R f' x $ i) at x within T" apply(unfold has_derivative_def, safe) apply(force simp: bounded_linear_def bounded_linear_axioms_def) using frechet_vec_nth[of x T f] assms unfolding has_derivative_def by auto lemma has_vderiv_on_vec_eq[simp]: fixes x::"real \ ('a::banach)^('n::finite)" shows "(D x = x' on T) = (\i. D (\t. x t $ i) = (\t. x' t $ i) on T)" unfolding has_vderiv_on_def has_vector_derivative_def apply safe using has_derivative_vec_nth has_derivative_vec_lambda by blast+ end \ No newline at end of file diff --git a/thys/Stirling_Formula/Ln_Gamma_Asymptotics.thy b/thys/Stirling_Formula/Ln_Gamma_Asymptotics.thy --- a/thys/Stirling_Formula/Ln_Gamma_Asymptotics.thy +++ b/thys/Stirling_Formula/Ln_Gamma_Asymptotics.thy @@ -1,1328 +1,1328 @@ (* File: Ln_Gamma_Asymptotics.thy Author: Manuel Eberl The complete asymptotics of the logarithmic Gamma function and the Polygamma functions. *) section \Complete asymptotics of the logarithmic Gamma function\ theory Ln_Gamma_Asymptotics imports - "HOL-Analysis.Analysis" + "HOL-Complex_Analysis.Complex_Analysis" Bernoulli.Bernoulli_FPS Bernoulli.Periodic_Bernpoly Stirling_Formula begin subsection \Auxiliary Facts\ (* TODO Move *) lemma filterlim_at_infinity_conv_norm_at_top: "filterlim f at_infinity G \ filterlim (\x. norm (f x)) at_top G" by (auto simp: filterlim_at_infinity[OF order.refl] filterlim_at_top_gt[of _ _ 0]) corollary Ln_times_of_nat: "\r > 0; z \ 0\ \ Ln(of_nat r * z :: complex) = ln (of_nat r) + Ln(z)" using Ln_times_of_real[of "of_nat r" z] by simp lemma tendsto_of_real_0_I: "(f \ 0) G \ ((\x. (of_real (f x))) \ (0 ::'a::real_normed_div_algebra)) G" by (subst (asm) tendsto_of_real_iff [symmetric]) simp lemma negligible_atLeastAtMostI: "b \ a \ negligible {a..(b::real)}" by (cases "b < a") auto lemma integrable_on_negligible: "negligible A \ (f :: 'n :: euclidean_space \ 'a :: banach) integrable_on A" by (subst integrable_spike_set_eq[of _ "{}"]) (simp_all add: integrable_on_empty) lemma vector_derivative_cong_eq: assumes "eventually (\x. x \ A \ f x = g x) (nhds x)" "x = y" "A = B" "x \ A" shows "vector_derivative f (at x within A) = vector_derivative g (at y within B)" proof - from eventually_nhds_x_imp_x[OF assms(1)] assms(4) have "f x = g x" by blast hence "(\D. (f has_vector_derivative D) (at x within A)) = (\D. (g has_vector_derivative D) (at x within A))" using assms by (intro ext has_vector_derivative_cong_ev refl assms) simp_all thus ?thesis by (simp add: vector_derivative_def assms) qed lemma differentiable_of_real [simp]: "of_real differentiable at x within A" proof - have "(of_real has_vector_derivative 1) (at x within A)" by (auto intro!: derivative_eq_intros) thus ?thesis by (rule differentiableI_vector) qed lemma higher_deriv_cong_ev: assumes "eventually (\x. f x = g x) (nhds x)" "x = y" shows "(deriv ^^ n) f x = (deriv ^^ n) g y" proof - from assms(1) have "eventually (\x. (deriv ^^ n) f x = (deriv ^^ n) g x) (nhds x)" proof (induction n arbitrary: f g) case (Suc n) from Suc.prems have "eventually (\y. eventually (\z. f z = g z) (nhds y)) (nhds x)" by (simp add: eventually_eventually) hence "eventually (\x. deriv f x = deriv g x) (nhds x)" by eventually_elim (rule deriv_cong_ev, simp_all) thus ?case by (auto intro!: deriv_cong_ev Suc simp: funpow_Suc_right simp del: funpow.simps) qed auto from eventually_nhds_x_imp_x[OF this] assms(2) show ?thesis by simp qed lemma deriv_of_real [simp]: "at x within A \ bot \ vector_derivative of_real (at x within A) = 1" by (auto intro!: vector_derivative_within derivative_eq_intros) lemma deriv_Re [simp]: "deriv Re = (\_. 1)" by (auto intro!: DERIV_imp_deriv simp: fun_eq_iff) lemma vector_derivative_of_real_left: assumes "f differentiable at x" shows "vector_derivative (\x. of_real (f x)) (at x) = of_real (deriv f x)" proof - have "vector_derivative (of_real \ f) (at x) = (of_real (deriv f x))" by (subst vector_derivative_chain_at) (simp_all add: scaleR_conv_of_real field_derivative_eq_vector_derivative assms) thus ?thesis by (simp add: o_def) qed lemma vector_derivative_of_real_right: assumes "f field_differentiable at (of_real x)" shows "vector_derivative (\x. f (of_real x)) (at x) = deriv f (of_real x)" proof - have "vector_derivative (f \ of_real) (at x) = deriv f (of_real x)" using assms by (subst vector_derivative_chain_at_general) simp_all thus ?thesis by (simp add: o_def) qed lemma Ln_holomorphic [holomorphic_intros]: assumes "A \ \\<^sub>\\<^sub>0 = {}" shows "Ln holomorphic_on (A :: complex set)" proof (intro holomorphic_onI) fix z assume "z \ A" with assms have "(Ln has_field_derivative inverse z) (at z within A)" by (auto intro!: derivative_eq_intros) thus "Ln field_differentiable at z within A" by (auto simp: field_differentiable_def) qed lemma ln_Gamma_holomorphic [holomorphic_intros]: assumes "A \ \\<^sub>\\<^sub>0 = {}" shows "ln_Gamma holomorphic_on (A :: complex set)" proof (intro holomorphic_onI) fix z assume "z \ A" with assms have "(ln_Gamma has_field_derivative Digamma z) (at z within A)" by (auto intro!: derivative_eq_intros) thus "ln_Gamma field_differentiable at z within A" by (auto simp: field_differentiable_def) qed lemma higher_deriv_Polygamma: assumes "z \ \\<^sub>\\<^sub>0" shows "(deriv ^^ n) (Polygamma m) z = Polygamma (m + n) (z :: 'a :: {real_normed_field,euclidean_space})" proof - have "eventually (\u. (deriv ^^ n) (Polygamma m) u = Polygamma (m + n) u) (nhds z)" proof (induction n) case (Suc n) from Suc.IH have "eventually (\z. eventually (\u. (deriv ^^ n) (Polygamma m) u = Polygamma (m + n) u) (nhds z)) (nhds z)" by (simp add: eventually_eventually) hence "eventually (\z. deriv ((deriv ^^ n) (Polygamma m)) z = deriv (Polygamma (m + n)) z) (nhds z)" by eventually_elim (intro deriv_cong_ev refl) moreover have "eventually (\z. z \ UNIV - \\<^sub>\\<^sub>0) (nhds z)" using assms by (intro eventually_nhds_in_open open_Diff open_UNIV) auto ultimately show ?case by eventually_elim (simp_all add: deriv_Polygamma) qed simp_all thus ?thesis by (rule eventually_nhds_x_imp_x) qed lemma higher_deriv_cmult: assumes "f holomorphic_on A" "x \ A" "open A" shows "(deriv ^^ j) (\x. c * f x) x = c * (deriv ^^ j) f x" using assms proof (induction j arbitrary: f x) case (Suc j f x) have "deriv ((deriv ^^ j) (\x. c * f x)) x = deriv (\x. c * (deriv ^^ j) f x) x" using eventually_nhds_in_open[of A x] assms(2,3) Suc.prems by (intro deriv_cong_ev refl) (auto elim!: eventually_mono simp: Suc.IH) also have "\ = c * deriv ((deriv ^^ j) f) x" using Suc.prems assms(2,3) by (intro deriv_cmult holomorphic_on_imp_differentiable_at holomorphic_higher_deriv) auto finally show ?case by simp qed simp_all lemma higher_deriv_ln_Gamma_complex: assumes "(x::complex) \ \\<^sub>\\<^sub>0" shows "(deriv ^^ j) ln_Gamma x = (if j = 0 then ln_Gamma x else Polygamma (j - 1) x)" proof (cases j) case (Suc j') have "(deriv ^^ j') (deriv ln_Gamma) x = (deriv ^^ j') Digamma x" using eventually_nhds_in_open[of "UNIV - \\<^sub>\\<^sub>0" x] assms by (intro higher_deriv_cong_ev refl) (auto elim!: eventually_mono simp: open_Diff deriv_ln_Gamma_complex) also have "\ = Polygamma j' x" using assms by (subst higher_deriv_Polygamma) (auto elim!: nonpos_Ints_cases simp: complex_nonpos_Reals_iff) finally show ?thesis using Suc by (simp del: funpow.simps add: funpow_Suc_right) qed simp_all lemma higher_deriv_ln_Gamma_real: assumes "(x::real) > 0" shows "(deriv ^^ j) ln_Gamma x = (if j = 0 then ln_Gamma x else Polygamma (j - 1) x)" proof (cases j) case (Suc j') have "(deriv ^^ j') (deriv ln_Gamma) x = (deriv ^^ j') Digamma x" using eventually_nhds_in_open[of "{0<..}" x] assms by (intro higher_deriv_cong_ev refl) (auto elim!: eventually_mono simp: open_Diff deriv_ln_Gamma_real) also have "\ = Polygamma j' x" using assms by (subst higher_deriv_Polygamma) (auto elim!: nonpos_Ints_cases simp: complex_nonpos_Reals_iff) finally show ?thesis using Suc by (simp del: funpow.simps add: funpow_Suc_right) qed simp_all lemma higher_deriv_ln_Gamma_complex_of_real: assumes "(x :: real) > 0" shows "(deriv ^^ j) ln_Gamma (complex_of_real x) = of_real ((deriv ^^ j) ln_Gamma x)" using assms by (auto simp: higher_deriv_ln_Gamma_real higher_deriv_ln_Gamma_complex ln_Gamma_complex_of_real Polygamma_of_real) (* END TODO *) lemma stirling_limit_aux1: "((\y. Ln (1 + z * of_real y) / of_real y) \ z) (at_right 0)" for z :: complex proof (cases "z = 0") case True then show ?thesis by simp next case False have "((\y. ln (1 + z * of_real y)) has_vector_derivative 1 * z) (at 0)" by (rule has_vector_derivative_real_field) (auto intro!: derivative_eq_intros) then have "(\y. (Ln (1 + z * of_real y) - of_real y * z) / of_real \y\) \0\ 0" by (auto simp add: has_vector_derivative_def has_derivative_def netlimit_at scaleR_conv_of_real field_simps) then have "((\y. (Ln (1 + z * of_real y) - of_real y * z) / of_real \y\) \ 0) (at_right 0)" by (rule filterlim_mono[OF _ _ at_le]) simp_all also have "?this \ ((\y. Ln (1 + z * of_real y) / (of_real y) - z) \ 0) (at_right 0)" using eventually_at_right_less[of "0::real"] by (intro filterlim_cong refl) (auto elim!: eventually_mono simp: field_simps) finally show ?thesis by (simp only: LIM_zero_iff) qed lemma stirling_limit_aux2: "((\y. y * Ln (1 + z / of_real y)) \ z) at_top" for z :: complex using stirling_limit_aux1[of z] by (subst filterlim_at_top_to_right) (simp add: field_simps) lemma Union_atLeastAtMost: assumes "N > 0" shows "(\n\{0.. {0..real N}" thus "x \ (\n\{0.. 0" "x < real N" by simp_all hence "x \ real (nat \x\)" "x \ real (nat \x\ + 1)" by linarith+ moreover from x have "nat \x\ < N" by linarith ultimately have "\n\{0.. {real n..real (n + 1)}" by (intro bexI[of _ "nat \x\"]) simp_all thus ?thesis by blast qed qed auto subsection \Asymptotics of @{term ln_Gamma}\ text \ This is the error term that occurs in the expansion of @{term ln_Gamma}. It can be shown to be of order $O(s^{-n})$. \ definition stirling_integral :: "nat \ 'a :: {real_normed_div_algebra, banach} \ 'a" where "stirling_integral n s = lim (\N. integral {0..N} (\x. of_real (pbernpoly n x) / (of_real x + s) ^ n))" text \ \ context fixes s :: complex assumes s: "Re s > 0" fixes approx :: "nat \ complex" defines "approx \ (\N. (\n = 1.. \\\ ln_Gamma s\\ (ln_Gamma (of_nat N) - ln (2 * pi / of_nat N) / 2 - of_nat N * ln (of_nat N) + of_nat N) - \ \\\ 0\\ s * (harm (N - 1) - ln (of_nat (N - 1)) - euler_mascheroni) + \ \\\ 0\\ s * (ln (of_nat N + s) - ln (of_nat (N - 1))) - \ \\\ 0\\ (1/2) * (ln (of_nat N + s) - ln (of_nat N)) + \ \\\ 0\\ of_nat N * (ln (of_nat N + s) - ln (of_nat N)) - \ \\\ s\\ (s - 1/2) * ln s - ln (2 * pi) / 2)" begin qualified lemma assumes N: "N > 0" shows integrable_pbernpoly_1: "(\x. of_real (-pbernpoly 1 x) / (of_real x + s)) integrable_on {0..real N}" and integral_pbernpoly_1_aux: "integral {0..real N} (\x. -of_real (pbernpoly 1 x) / (of_real x + s)) = approx N" and has_integral_pbernpoly_1: "((\x. pbernpoly 1 x /(x + s)) has_integral (\mx. -pbernpoly 1 x / (x + s)) has_integral (of_nat n + 1/2 + s) * (ln (of_nat (n + 1) + s) - ln (of_nat n + s)) - 1) {of_nat n..of_nat (n + 1)}" for n proof (rule has_integral_spike) have "((\x. (of_nat n + 1/2 + s) * (1 / (x + s)) - 1) has_integral (of_nat n + 1/2 + s) * (ln (of_real (real (n + 1)) + s) - ln (of_real (real n) + s)) - 1) {of_nat n..of_nat (n + 1)}" using s has_integral_const_real[of 1 "of_nat n" "of_nat (n + 1)"] by (intro has_integral_diff has_integral_mult_right fundamental_theorem_of_calculus) (auto intro!: derivative_eq_intros has_vector_derivative_real_field simp: has_field_derivative_iff_has_vector_derivative [symmetric] field_simps complex_nonpos_Reals_iff) thus "((\x. (of_nat n + 1/2 + s) * (1 / (x + s)) - 1) has_integral (of_nat n + 1/2 + s) * (ln (of_nat (n + 1) + s) - ln (of_nat n + s)) - 1) {of_nat n..of_nat (n + 1)}" by simp show "-pbernpoly 1 x / (x + s) = (of_nat n + 1/2 + s) * (1 / (x + s)) - 1" if "x \ {of_nat n..of_nat (n + 1)} - {of_nat (n + 1)}" for x proof - have x: "x \ real n" "x < real (n + 1)" using that by simp_all hence "floor x = int n" by linarith moreover from s x have "complex_of_real x \ -s" by (auto simp add: complex_eq_iff simp del: of_nat_Suc) ultimately show "-pbernpoly 1 x / (x + s) = (of_nat n + 1/2 + s) * (1 / (x + s)) - 1" by (auto simp: pbernpoly_def bernpoly_def frac_def divide_simps add_eq_0_iff2) qed qed simp_all hence *: "\I. I\?A \ ((\x. -pbernpoly 1 x / (x + s)) has_integral (Inf I + 1/2 + s) * (ln (Inf I + 1 + s) - ln (Inf I + s)) - 1) I" by (auto simp: add_ac) have "((\x. - pbernpoly 1 x / (x + s)) has_integral (\I\?A. (Inf I + 1 / 2 + s) * (ln (Inf I + 1 + s) - ln (Inf I + s)) - 1)) (\n\{0..x. - pbernpoly 1 x / (x + s)) has_integral ?i) {0..real N}" by (subst has_integral_spike_set_eq) (use Union_atLeastAtMost assms in \auto simp: intro!: empty_imp_negligible\) hence "(\x. - pbernpoly 1 x / (x + s)) integrable_on {0..real N}" and integral: "integral {0..real N} (\x. - pbernpoly 1 x / (x + s)) = ?i" by (simp_all add: has_integral_iff) show "(\x. - pbernpoly 1 x / (x + s)) integrable_on {0..real N}" by fact note has_integral_neg[OF has_integral] also have "-?i = (\xx. of_real (pbernpoly 1 x) / (of_real x + s)) has_integral \) {0..real N}" by simp note integral also have "?i = (\nnnn=1.. = (\n=1..n=1..n=1..n=1.. - (\n = 1..n=1.. 0" for x by (auto simp: complex_eq_iff) hence "(\n=1..n=1.. = ln (fact (N - 1)) + (\n=1..n=1..n=1..n=1..n = 1..n = 1..x. -of_real (pbernpoly 1 x) / (of_real x + s)) = \" by simp qed lemma integrable_ln_Gamma_aux: shows "(\x. of_real (pbernpoly n x) / (of_real x + s) ^ n) integrable_on {0..real N}" proof (cases "n = 1") case True with s show ?thesis using integrable_neg[OF integrable_pbernpoly_1[of N]] by (cases "N = 0") (simp_all add: integrable_on_negligible) next case False from s have "of_real x + s \ 0" if "x \ 0" for x using that by (auto simp: complex_eq_iff add_eq_0_iff2) with False s show ?thesis by (auto intro!: integrable_continuous_real continuous_intros) qed text \ This following proof is based on ``Rudiments of the theory of the gamma function'' by Bruce Berndt~\cite{berndt}. \ qualified lemma integral_pbernpoly_1: "(\N. integral {0..real N} (\x. pbernpoly 1 x / (x + s))) \ -ln_Gamma s - s + (s - 1 / 2) * ln s + ln (2 * pi) / 2" proof - have neq: "s + of_real x \ 0" if "x \ 0" for x :: real using that s by (auto simp: complex_eq_iff) have "(approx \ ln_Gamma s - 0 - 0 + 0 - 0 + s - (s - 1/2) * ln s - ln (2 * pi) / 2) at_top" unfolding approx_def proof (intro tendsto_add tendsto_diff) from s have s': "s \ \\<^sub>\\<^sub>0" by (auto elim!: nonpos_Ints_cases) have "(\n. \i=1.. ln_Gamma s + euler_mascheroni * s + ln s" (is "?f \ _") using ln_Gamma_series'_aux[OF s'] unfolding sums_def by (subst LIMSEQ_Suc_iff [symmetric], subst (asm) sum.atLeast1_atMost_eq [symmetric]) (simp add: atLeastLessThanSuc_atLeastAtMost) thus "((\n. ?f n - (euler_mascheroni * s + ln s)) \ ln_Gamma s) at_top" by (auto intro: tendsto_eq_intros) next show "(\x. complex_of_real (ln_Gamma (real x) - ln (2 * pi / real x) / 2 - real x * ln (real x) + real x)) \ 0" proof (intro tendsto_of_real_0_I filterlim_compose[OF tendsto_sandwich filterlim_real_sequentially]) show "eventually (\x::real. ln_Gamma x - ln (2 * pi / x) / 2 - x * ln x + x \ 0) at_top" using eventually_ge_at_top[of "1::real"] by eventually_elim (insert ln_Gamma_bounds(1), simp add: algebra_simps) show "eventually (\x::real. ln_Gamma x - ln (2 * pi / x) / 2 - x * ln x + x \ 1 / 12 * inverse x) at_top" using eventually_ge_at_top[of "1::real"] by eventually_elim (insert ln_Gamma_bounds(2), simp add: field_simps) show "((\x::real. 1 / 12 * inverse x) \ 0) at_top" by (intro tendsto_mult_right_zero tendsto_inverse_0_at_top filterlim_ident) qed simp_all next have "(\x. s * of_real (harm (x - 1) - ln (real (x - 1)) - euler_mascheroni)) \ s * of_real (euler_mascheroni - euler_mascheroni)" by (subst LIMSEQ_Suc_iff [symmetric], intro tendsto_intros) (insert euler_mascheroni_LIMSEQ, simp_all) also have "?this \ (\x. s * (harm (x - 1) - ln (of_nat (x - 1)) - euler_mascheroni)) \ 0" by (intro filterlim_cong refl eventually_mono[OF eventually_gt_at_top[of "1::nat"]]) (auto simp: Ln_of_nat of_real_harm) finally show "(\x. s * (harm (x - 1) - ln (of_nat (x - 1)) - euler_mascheroni)) \ 0" . next have "((\x. ln (1 + (s + 1) / of_real x)) \ ln (1 + 0)) at_top" (is ?P) by (intro tendsto_intros tendsto_divide_0[OF tendsto_const]) (simp_all add: filterlim_ident filterlim_at_infinity_conv_norm_at_top filterlim_abs_real) also have "ln (of_real (x + 1) + s) - ln (complex_of_real x) = ln (1 + (s + 1) / of_real x)" if "x > 1" for x using that s using Ln_divide_of_real[of x "of_real (x + 1) + s", symmetric] neq[of "x+1"] by (simp add: field_simps Ln_of_real) hence "?P \ ((\x. ln (of_real (x + 1) + s) - ln (of_real x)) \ 0) at_top" by (intro filterlim_cong refl) (auto intro: eventually_mono[OF eventually_gt_at_top[of "1::real"]]) finally have "((\n. ln (of_real (real n + 1) + s) - ln (of_real (real n))) \ 0) at_top" by (rule filterlim_compose[OF _ filterlim_real_sequentially]) hence "((\n. ln (of_nat n + s) - ln (of_nat (n - 1))) \ 0) at_top" by (subst LIMSEQ_Suc_iff [symmetric]) (simp add: add_ac) thus "(\x. s * (ln (of_nat x + s) - ln (of_nat (x - 1)))) \ 0" by (rule tendsto_mult_right_zero) next have "((\x. ln (1 + s / of_real x)) \ ln (1 + 0)) at_top" (is ?P) by (intro tendsto_intros tendsto_divide_0[OF tendsto_const]) (simp_all add: filterlim_ident filterlim_at_infinity_conv_norm_at_top filterlim_abs_real) also have "ln (of_real x + s) - ln (of_real x) = ln (1 + s / of_real x)" if "x > 0" for x using Ln_divide_of_real[of x "of_real x + s"] neq[of x] that by (auto simp: field_simps Ln_of_real) hence "?P \ ((\x. ln (of_real x + s) - ln (of_real x)) \ 0) at_top" using s by (intro filterlim_cong refl) (auto intro: eventually_mono [OF eventually_gt_at_top[of "1::real"]]) finally have "(\x. (1/2) * (ln (of_real (real x) + s) - ln (of_real (real x)))) \ 0" by (rule tendsto_mult_right_zero[OF filterlim_compose[OF _ filterlim_real_sequentially]]) thus "(\x. (1/2) * (ln (of_nat x + s) - ln (of_nat x))) \ 0" by simp next have "((\x. x * (ln (1 + s / of_real x))) \ s) at_top" (is ?P) by (rule stirling_limit_aux2) also have "ln (1 + s / of_real x) = ln (of_real x + s) - ln (of_real x)" if "x > 1" for x using that s Ln_divide_of_real [of x "of_real x + s", symmetric] neq[of x] by (auto simp: Ln_of_real field_simps) hence "?P \ ((\x. of_real x * (ln (of_real x + s) - ln (of_real x))) \ s) at_top" by (intro filterlim_cong refl) (auto intro: eventually_mono[OF eventually_gt_at_top[of "1::real"]]) finally have "(\n. of_real (real n) * (ln (of_real (real n) + s) - ln (of_real (real n)))) \ s" by (rule filterlim_compose[OF _ filterlim_real_sequentially]) thus "(\n. of_nat n * (ln (of_nat n + s) - ln (of_nat n))) \ s" by simp qed simp_all also have "?this \ ((\N. integral {0..real N} (\x. -pbernpoly 1 x / (x + s))) \ ln_Gamma s + s - (s - 1/2) * ln s - ln (2 * pi) / 2) at_top" using integral_pbernpoly_1_aux by (intro filterlim_cong refl) (auto intro: eventually_mono[OF eventually_gt_at_top[of "0::nat"]]) also have "(\N. integral {0..real N} (\x. -pbernpoly 1 x / (x + s))) = (\N. -integral {0..real N} (\x. pbernpoly 1 x / (x + s)))" by (simp add: fun_eq_iff) finally show ?thesis by (simp add: tendsto_minus_cancel_left [symmetric] algebra_simps) qed qualified lemma pbernpoly_integral_conv_pbernpoly_integral_Suc: assumes "n \ 1" shows "integral {0..real N} (\x. pbernpoly n x / (x + s) ^ n) = of_real (pbernpoly (Suc n) (real N)) / (of_nat (Suc n) * (s + of_nat N) ^ n) - of_real (bernoulli (Suc n)) / (of_nat (Suc n) * s ^ n) + of_nat n / of_nat (Suc n) * integral {0..real N} (\x. of_real (pbernpoly (Suc n) x) / (of_real x + s) ^ Suc n)" proof - note [derivative_intros] = has_field_derivative_pbernpoly_Suc' define I where "I = -of_real (pbernpoly (Suc n) (of_nat N)) / (of_nat (Suc n) * (of_nat N + s) ^ n) + of_real (bernoulli (Suc n) / real (Suc n)) / s ^ n + integral {0..real N} (\x. of_real (pbernpoly n x) / (of_real x + s) ^ n)" have "((\x. (-of_nat n * inverse (of_real x + s) ^ Suc n) * (of_real (pbernpoly (Suc n) x) / (of_nat (Suc n)))) has_integral -I) {0..real N}" proof (rule integration_by_parts_interior_strong[OF bounded_bilinear_mult]) fix x :: real assume x: "x \ {0<.. \" proof assume "x \ \" then obtain n where "x = of_int n" by (auto elim!: Ints_cases) with x have x': "x = of_nat (nat n)" by simp from x show False by (auto simp: x') qed hence "((\x. of_real (pbernpoly (Suc n) x / of_nat (Suc n))) has_vector_derivative complex_of_real (pbernpoly n x)) (at x)" by (intro has_vector_derivative_of_real) (auto intro!: derivative_eq_intros) thus "((\x. of_real (pbernpoly (Suc n) x) / of_nat (Suc n)) has_vector_derivative complex_of_real (pbernpoly n x)) (at x)" by simp from x s have "complex_of_real x + s \ 0" by (auto simp: complex_eq_iff) thus "((\x. inverse (of_real x + s) ^ n) has_vector_derivative - of_nat n * inverse (of_real x + s) ^ Suc n) (at x)" using x s assms by (auto intro!: derivative_eq_intros has_vector_derivative_real_field simp: divide_simps power_add [symmetric] simp del: power_Suc) next have "complex_of_real x + s \ 0" if "x \ 0" for x using that s by (auto simp: complex_eq_iff) thus "continuous_on {0..real N} (\x. inverse (of_real x + s) ^ n)" "continuous_on {0..real N} (\x. complex_of_real (pbernpoly (Suc n) x) / of_nat (Suc n))" using assms s by (auto intro!: continuous_intros simp del: of_nat_Suc) next have "((\x. inverse (of_real x + s) ^ n * of_real (pbernpoly n x)) has_integral pbernpoly (Suc n) (of_nat N) / (of_nat (Suc n) * (of_nat N + s) ^ n) - of_real (bernoulli (Suc n) / real (Suc n)) / s ^ n - -I) {0..real N}" using integrable_ln_Gamma_aux[of n N] assms by (auto simp: I_def has_integral_integral divide_simps) thus "((\x. inverse (of_real x + s) ^ n * of_real (pbernpoly n x)) has_integral inverse (of_real (real N) + s) ^ n * (of_real (pbernpoly (Suc n) (real N)) / of_nat (Suc n)) - inverse (of_real 0 + s) ^ n * (of_real (pbernpoly (Suc n) 0) / of_nat (Suc n)) - - I) {0..real N}" by (simp_all add: field_simps) qed simp_all also have "(\x. - of_nat n * inverse (of_real x + s) ^ Suc n * (of_real (pbernpoly (Suc n) x) / of_nat (Suc n))) = (\x. - (of_nat n / of_nat (Suc n) * of_real (pbernpoly (Suc n) x) / (of_real x + s) ^ Suc n))" by (simp add: divide_simps fun_eq_iff) finally have "((\x. - (of_nat n / of_nat (Suc n) * of_real (pbernpoly (Suc n) x) / (of_real x + s) ^ Suc n)) has_integral - I) {0..real N}" . from has_integral_neg[OF this] show ?thesis by (auto simp add: I_def has_integral_iff algebra_simps integral_mult_right [symmetric] simp del: power_Suc of_nat_Suc ) qed lemma pbernpoly_over_power_tendsto_0: assumes "n > 0" shows "(\x. of_real (pbernpoly (Suc n) (real x)) / (of_nat (Suc n) * (s + of_nat x) ^ n)) \ 0" proof - from s have neq: "s + of_nat n \ 0" for n by (auto simp: complex_eq_iff) from bounded_pbernpoly[of "Suc n"] guess c . note c = this have "eventually (\x. norm (of_real (pbernpoly (Suc n) (real x)) / (of_nat (Suc n) * (s + of_nat x) ^ n)) \ (c / real (Suc n)) / real x ^ n) at_top" using eventually_gt_at_top[of "0::nat"] proof eventually_elim case (elim x) have "norm (of_real (pbernpoly (Suc n) (real x)) / (of_nat (Suc n) * (s + of_nat x) ^ n)) \ (c / real (Suc n)) / norm (s + of_nat x) ^ n" (is "_ \ ?rhs") using c[of x] by (auto simp: norm_divide norm_mult norm_power neq field_simps simp del: of_nat_Suc) also have "real x \ cmod (s + of_nat x)" using complex_Re_le_cmod[of "s + of_nat x"] s by simp hence "?rhs \ (c / real (Suc n)) / real x ^ n" using s elim c[of 0] neq[of x] by (intro divide_left_mono power_mono mult_pos_pos divide_nonneg_pos zero_less_power) auto finally show ?case . qed moreover have "(\x. (c / real (Suc n)) / real x ^ n) \ 0" by (intro real_tendsto_divide_at_top[OF tendsto_const] filterlim_pow_at_top assms filterlim_real_sequentially) ultimately show ?thesis by (rule Lim_null_comparison) qed lemma convergent_stirling_integral: assumes "n > 0" shows "convergent (\N. integral {0..real N} (\x. of_real (pbernpoly n x) / (of_real x + s) ^ n))" (is "convergent (?f n)") proof - have "convergent (?f (Suc n))" for n proof (induction n) case 0 thus ?case using integral_pbernpoly_1 by (auto intro!: convergentI) next case (Suc n) have "convergent (\N. ?f (Suc n) N - of_real (pbernpoly (Suc (Suc n)) (real N)) / (of_nat (Suc (Suc n)) * (s + of_nat N) ^ Suc n) + of_real (bernoulli (Suc (Suc n)) / (real (Suc (Suc n)))) / s ^ Suc n)" (is "convergent ?g") by (intro convergent_add convergent_diff Suc convergent_const convergentI[OF pbernpoly_over_power_tendsto_0]) simp_all also have "?g = (\N. of_nat (Suc n) / of_nat (Suc (Suc n)) * ?f (Suc (Suc n)) N)" using s by (subst pbernpoly_integral_conv_pbernpoly_integral_Suc) (auto simp: fun_eq_iff field_simps simp del: of_nat_Suc power_Suc) also have "convergent \ \ convergent (?f (Suc (Suc n)))" by (intro convergent_mult_const_iff) (simp_all del: of_nat_Suc) finally show ?case . qed from this[of "n - 1"] assms show ?thesis by simp qed lemma stirling_integral_conv_stirling_integral_Suc: assumes "n > 0" shows "stirling_integral n s = of_nat n / of_nat (Suc n) * stirling_integral (Suc n) s - of_real (bernoulli (Suc n)) / (of_nat (Suc n) * s ^ n)" proof - have "(\N. of_real (pbernpoly (Suc n) (real N)) / (of_nat (Suc n) * (s + of_nat N) ^ n) - of_real (bernoulli (Suc n)) / (real (Suc n) * s ^ n) + integral {0..real N} (\x. of_nat n / of_nat (Suc n) * (of_real (pbernpoly (Suc n) x) / (of_real x + s) ^ Suc n))) \ 0 - of_real (bernoulli (Suc n)) / (of_nat (Suc n) * s ^ n) + of_nat n / of_nat (Suc n) * stirling_integral (Suc n) s" (is "?f \ _") unfolding stirling_integral_def integral_mult_right using convergent_stirling_integral[of "Suc n"] assms s by (intro tendsto_intros pbernpoly_over_power_tendsto_0) (auto simp: convergent_LIMSEQ_iff simp del: of_nat_Suc) also have "?this \ (\N. integral {0..real N} (\x. of_real (pbernpoly n x) / (of_real x + s) ^ n)) \ of_nat n / of_nat (Suc n) * stirling_integral (Suc n) s - of_real (bernoulli (Suc n)) / (of_nat (Suc n) * s ^ n)" using eventually_gt_at_top[of "0::nat"] pbernpoly_integral_conv_pbernpoly_integral_Suc[of n] assms unfolding integral_mult_right by (intro filterlim_cong refl) (auto elim!: eventually_mono simp del: power_Suc) finally show ?thesis unfolding stirling_integral_def[of n] by (rule limI) qed lemma stirling_integral_1_unfold: assumes "m > 0" shows "stirling_integral 1 s = stirling_integral m s / of_nat m - (\k=1..k=1..k = 1.. 0" shows "ln_Gamma s = (s - 1 / 2) * ln s - s + ln (2 * pi) / 2 + (\k=1..k = 1.. 0 \ (\x. integral {0..real x} (\x. of_real (pbernpoly n x) / (of_real x + s) ^ n)) \ stirling_integral n s" unfolding stirling_integral_def using convergent_stirling_integral[of n] by (simp only: convergent_LIMSEQ_iff) end lemmas has_integral_of_real = has_integral_linear[OF _ bounded_linear_of_real, unfolded o_def] lemmas integral_of_real = integral_linear[OF _ bounded_linear_of_real, unfolded o_def] lemma integrable_ln_Gamma_aux_real: assumes "0 < s" shows "(\x. pbernpoly n x / (x + s) ^ n) integrable_on {0..real N}" proof - have "(\x. complex_of_real (pbernpoly n x / (x + s) ^ n)) integrable_on {0..real N}" using integrable_ln_Gamma_aux[of "of_real s" n N] assms by simp from integrable_linear[OF this bounded_linear_Re] show ?thesis by (simp only: o_def Re_complex_of_real) qed lemma assumes "x > 0" "n > 0" shows stirling_integral_complex_of_real: "stirling_integral n (complex_of_real x) = of_real (stirling_integral n x)" and LIMSEQ_stirling_integral_real: "(\N. integral {0..real N} (\t. pbernpoly n t / (t + x) ^ n)) \ stirling_integral n x" and stirling_integral_real_convergent: "convergent (\N. integral {0..real N} (\t. pbernpoly n t / (t + x) ^ n))" proof - have "(\N. integral {0..real N} (\t. of_real (pbernpoly n t / (t + x) ^ n))) \ stirling_integral n (complex_of_real x)" using LIMSEQ_stirling_integral[of "complex_of_real x" n] assms by simp hence "(\N. of_real (integral {0..real N} (\t. pbernpoly n t / (t + x) ^ n))) \ stirling_integral n (complex_of_real x)" using integrable_ln_Gamma_aux_real[OF assms(1), of n] by (subst (asm) integral_of_real) simp from tendsto_Re[OF this] have "(\N. integral {0..real N} (\t. pbernpoly n t / (t + x) ^ n)) \ Re (stirling_integral n (complex_of_real x))" by simp thus "convergent (\N. integral {0..real N} (\t. pbernpoly n t / (t + x) ^ n))" by (rule convergentI) thus "(\N. integral {0..real N} (\t. pbernpoly n t / (t + x) ^ n)) \ stirling_integral n x" unfolding stirling_integral_def by (simp add: convergent_LIMSEQ_iff) from tendsto_of_real[OF this, where 'a = complex] integrable_ln_Gamma_aux_real[OF assms(1), of n] have "(\xa. integral {0..real xa} (\xa. complex_of_real (pbernpoly n xa) / (complex_of_real xa + x) ^ n)) \ complex_of_real (stirling_integral n x)" by (subst (asm) integral_of_real [symmetric]) simp_all from LIMSEQ_unique[OF this LIMSEQ_stirling_integral[of "complex_of_real x" n]] assms show "stirling_integral n (complex_of_real x) = of_real (stirling_integral n x)" by simp qed lemma ln_Gamma_stirling_real: assumes "x > (0 :: real)" "m > (0::nat)" shows "ln_Gamma x = (x - 1 / 2) * ln x - x + ln (2 * pi) / 2 + (\k = 1..k = 1.. (1::nat)" obtains c where "\s. Re s > 0 \ norm (stirling_integral n s) \ c / Re s ^ (n - 1)" proof - obtain c where c: "norm (pbernpoly n x) \ c" for x by (rule bounded_pbernpoly[of n]) blast have c': "pbernpoly n x \ c" for x using c[of x] by (simp add: abs_real_def split: if_splits) from c[of 0] have c_nonneg: "c \ 0" by simp have "norm (stirling_integral n s) \ c / (real n - 1) / Re s ^ (n - 1)" if s: "Re s > 0" for s proof (rule Lim_norm_ubound[OF _ LIMSEQ_stirling_integral]) have pos: "x + norm s > 0" if "x \ 0" for x using s that by (intro add_nonneg_pos) auto have nz: "of_real x + s \ 0" if "x \ 0" for x using s that by (auto simp: complex_eq_iff) let ?bound = "\N. c / (Re s ^ (n - 1) * (real n - 1)) - c / ((real N + Re s) ^ (n - 1) * (real n - 1))" show "eventually (\N. norm (integral {0..real N} (\x. of_real (pbernpoly n x) / (of_real x + s) ^ n)) \ c / (real n - 1) / Re s ^ (n - 1)) at_top" using eventually_gt_at_top[of "0::nat"] proof eventually_elim case (elim N) let ?F = "\x. -c / ((x + Re s) ^ (n - 1) * (real n - 1))" from n s have "((\x. c / (x + Re s) ^ n) has_integral (?F (real N) - ?F 0)) {0..real N}" by (intro fundamental_theorem_of_calculus) (auto intro!: derivative_eq_intros simp: divide_simps power_diff add_eq_0_iff2 has_field_derivative_iff_has_vector_derivative [symmetric]) also have "?F (real N) - ?F 0 = ?bound N" by simp finally have *: "((\x. c / (x + Re s) ^ n) has_integral ?bound N) {0..real N}" . have "norm (integral {0..real N} (\x. of_real (pbernpoly n x) / (of_real x + s) ^ n)) \ integral {0..real N} (\x. c / (x + Re s) ^ n)" proof (intro integral_norm_bound_integral integrable_ln_Gamma_aux s ballI) fix x assume x: "x \ {0..real N}" have "norm (of_real (pbernpoly n x) / (of_real x + s) ^ n) \ c / norm (of_real x + s) ^ n" unfolding norm_divide norm_power using c by (intro divide_right_mono) simp_all also have "\ \ c / (x + Re s) ^ n" using x c c_nonneg s nz[of x] complex_Re_le_cmod[of "of_real x + s"] by (intro divide_left_mono power_mono mult_pos_pos zero_less_power add_nonneg_pos) auto finally show "norm (of_real (pbernpoly n x) / (of_real x + s) ^ n) \ \" . qed (insert n s * pos nz c, auto) also have "\ = ?bound N" using * by (simp add: has_integral_iff) also have "\ \ c / (Re s ^ (n - 1) * (real n - 1))" using c_nonneg elim s n by simp also have "\ = c / (real n - 1) / (Re s ^ (n - 1))" by simp finally show "norm (integral {0..real N} (\x. of_real (pbernpoly n x) / (of_real x + s) ^ n)) \ c / (real n - 1) / Re s ^ (n - 1)" . qed qed (insert s n, simp_all) thus ?thesis by (rule that) qed lemma stirling_integral_bound: assumes "n > 0" obtains c where "\s. Re s > 0 \ norm (stirling_integral n s) \ c / Re s ^ n" proof - let ?f = "\s. of_nat n / of_nat (Suc n) * stirling_integral (Suc n) s - of_real (bernoulli (Suc n)) / (of_nat (Suc n) * s ^ n)" from stirling_integral_bound_aux[of "Suc n"] assms obtain c where c: "\s. Re s > 0 \ norm (stirling_integral (Suc n) s) \ c / Re s ^ n" by auto define c1 where "c1 = real n / real (Suc n) * c" define c2 where "c2 = \bernoulli (Suc n)\ / real (Suc n)" have c2_nonneg: "c2 \ 0" by (simp add: c2_def) show ?thesis proof (rule that) fix s :: complex assume s: "Re s > 0" have "stirling_integral n s = ?f s" using s assms by (rule stirling_integral_conv_stirling_integral_Suc) also have "norm \ \ norm (of_nat n / of_nat (Suc n) * stirling_integral (Suc n) s) + norm (of_real (bernoulli (Suc n)) / (of_nat (Suc n) * s ^ n))" by (rule norm_triangle_ineq4) also have "\ = real n / real (Suc n) * norm (stirling_integral (Suc n) s) + c2 / norm s ^ n" (is "_ = ?A + ?B") by (simp add: norm_divide norm_mult norm_power c2_def field_simps del: of_nat_Suc) also have "?A \ real n / real (Suc n) * (c / Re s ^ n)" by (intro mult_left_mono c s) simp_all also have "\ = c1 / Re s ^ n" by (simp add: c1_def) also have "c2 / norm s ^ n \ c2 / Re s ^ n" using s c2_nonneg by (intro divide_left_mono power_mono complex_Re_le_cmod mult_pos_pos zero_less_power) auto also have "c1 / Re s ^ n + c2 / Re s ^ n = (c1 + c2) / Re s ^ n" using s by (simp add: field_simps) finally show "norm (stirling_integral n s) \ (c1 + c2) / Re s ^ n" by - simp_all qed qed end lemma stirling_integral_holomorphic [holomorphic_intros]: assumes m: "m > 0" and "\s\A. Re s > 0" shows "stirling_integral m holomorphic_on A" proof - let ?f = "\s::complex. of_nat m * ((s - 1 / 2) * Ln s - s + of_real (ln (2 * pi) / 2) + (\k=1.. stirling_integral m holomorphic_on A" using assms by (intro holomorphic_cong refl) (simp_all add: field_simps ln_Gamma_stirling_complex) finally show "stirling_integral m holomorphic_on A" . qed lemma stirling_integral_continuous_on [continuous_intros]: assumes m: "m > 0" and "\s\A. Re s > 0" shows "continuous_on A (stirling_integral m)" by (intro holomorphic_on_imp_continuous_on stirling_integral_holomorphic assms) lemma has_field_derivative_stirling_integral: assumes "Re x > 0" "n > 0" shows "(stirling_integral n has_field_derivative deriv (stirling_integral n) x) (at x)" using assms by (intro holomorphic_derivI[OF stirling_integral_holomorphic, of n "{s. Re s > 0}"] open_halfspace_Re_gt) auto lemma assumes n: "n > 0" and "x > 0" shows deriv_stirling_integral_complex_of_real: "(deriv ^^ j) (stirling_integral n) (complex_of_real x) = complex_of_real ((deriv ^^ j) (stirling_integral n) x)" (is "?lhs x = ?rhs x") and differentiable_stirling_integral_real: "(deriv ^^ j) (stirling_integral n) field_differentiable at x" (is ?thesis2) proof - let ?A = "{s. Re s > 0}" let ?f = "\j x. (deriv ^^ j) (stirling_integral n) (complex_of_real x)" let ?f' = "\j x. complex_of_real ((deriv ^^ j) (stirling_integral n) x)" have [simp]: "open ?A" by (simp add: open_halfspace_Re_gt) have "?lhs x = ?rhs x \ (deriv ^^ j) (stirling_integral n) field_differentiable at x" if "x > 0" for x using that proof (induction j arbitrary: x) case 0 have "((\x. Re (stirling_integral n (of_real x))) has_field_derivative Re (deriv (\x. stirling_integral n x) (of_real x))) (at x)" using 0 n by (auto intro!: derivative_intros has_vector_derivative_real_field field_differentiable_derivI holomorphic_on_imp_differentiable_at[of _ ?A] stirling_integral_holomorphic) also have "?this \ (stirling_integral n has_field_derivative Re (deriv (\x. stirling_integral n x) (of_real x))) (at x)" using eventually_nhds_in_open[of "{0<..}" x] 0 n by (intro has_field_derivative_cong_ev refl) (auto elim!: eventually_mono simp: stirling_integral_complex_of_real) finally have "stirling_integral n field_differentiable at x" by (auto simp: field_differentiable_def) with 0 n show ?case by (auto simp: stirling_integral_complex_of_real) next case (Suc j x) note IH = conjunct1[OF Suc.IH] conjunct2[OF Suc.IH] have *: "(deriv ^^ Suc j) (stirling_integral n) (complex_of_real x) = of_real ((deriv ^^ Suc j) (stirling_integral n) x)" if x: "x > 0" for x proof - have "deriv ((deriv ^^ j) (stirling_integral n)) (complex_of_real x) = vector_derivative (\x. (deriv ^^ j) (stirling_integral n) (of_real x)) (at x)" using n x by (intro vector_derivative_of_real_right [symmetric] holomorphic_on_imp_differentiable_at[of _ ?A] holomorphic_higher_deriv stirling_integral_holomorphic) auto also have "\ = vector_derivative (\x. of_real ((deriv ^^ j) (stirling_integral n) x)) (at x)" using eventually_nhds_in_open[of "{0<..}" x] x by (intro vector_derivative_cong_eq) (auto elim!: eventually_mono simp: IH(1)) also have "\ = of_real (deriv ((deriv ^^ j) (stirling_integral n)) x)" by (intro vector_derivative_of_real_left holomorphic_on_imp_differentiable_at[of _ ?A] field_differentiable_imp_differentiable IH(2) x) finally show ?thesis by simp qed have "((\x. Re ((deriv ^^ Suc j) (stirling_integral n) (of_real x))) has_field_derivative Re (deriv ((deriv ^^ Suc j) (stirling_integral n)) (of_real x))) (at x)" using Suc.prems n by (intro derivative_intros has_vector_derivative_real_field field_differentiable_derivI holomorphic_on_imp_differentiable_at[of _ ?A] stirling_integral_holomorphic holomorphic_higher_deriv) auto also have "?this \ ((deriv ^^ Suc j) (stirling_integral n) has_field_derivative Re (deriv ((deriv ^^ Suc j) (stirling_integral n)) (of_real x))) (at x)" using eventually_nhds_in_open[of "{0<..}" x] Suc.prems * by (intro has_field_derivative_cong_ev refl) (auto elim!: eventually_mono) finally have "(deriv ^^ Suc j) (stirling_integral n) field_differentiable at x" by (auto simp: field_differentiable_def) with *[OF Suc.prems] show ?case by blast qed from this[OF assms(2)] show "?lhs x = ?rhs x" ?thesis2 by blast+ qed text \ Unfortunately, asymptotic power series cannot, in general, be differentiated. However, since @{term ln_Gamma} is holomorphic on the entire positive real half-space, we can differentiate its asymptotic expansion after all. To do this, we use an ad-hoc version of the more general approach outlined in Erdelyi's ``Asymptotic Expansions'' for holomorphic functions: We bound the value of the $j$-th derivative of the remainder term at some value $x$ by applying Cauchy's integral formula along a circle centred at $x$ with radius $\frac{1}{2} x$. \ lemma deriv_stirling_integral_real_bound: assumes m: "m > 0" shows "(deriv ^^ j) (stirling_integral m) \ O(\x::real. 1 / x ^ (m + j))" proof - from stirling_integral_bound[OF m] guess c . note c = this have "0 \ cmod (stirling_integral m 1)" by simp also have "\ \ c" using c[of 1] by simp finally have c_nonneg: "c \ 0" . define B where "B = c * 2 ^ (m + Suc j)" define B' where "B' = B * fact j / 2" have "eventually (\x::real. norm ((deriv ^^ j) (stirling_integral m) x) \ B' * norm (1 / x ^ (m+ j))) at_top" using eventually_gt_at_top[of "0::real"] proof eventually_elim case (elim x) have "Re s > 0" if "s \ cball (of_real x) (x/2)" for s :: complex proof - have "x - Re s \ norm (of_real x - s)" using complex_Re_le_cmod[of "of_real x - s"] by simp also from that have "\ \ x/2" by (simp add: dist_complex_def) finally show ?thesis using elim by simp qed hence "((\u. stirling_integral m u / (u - of_real x) ^ Suc j) has_contour_integral complex_of_real (2 * pi) * \ / fact j * (deriv ^^ j) (stirling_integral m) (of_real x)) (circlepath (of_real x) (x/2))" using m elim by (intro Cauchy_has_contour_integral_higher_derivative_circlepath stirling_integral_continuous_on stirling_integral_holomorphic) auto hence "norm (of_real (2 * pi) * \ / fact j * (deriv ^^ j) (stirling_integral m) (of_real x)) \ B / x ^ (m + Suc j) * (2 * pi * (x / 2))" proof (rule has_contour_integral_bound_circlepath) fix u :: complex assume dist: "norm (u - of_real x) = x / 2" have "Re (of_real x - u) \ norm (of_real x - u)" by (rule complex_Re_le_cmod) also have "\ = x / 2" using dist by (simp add: norm_minus_commute) finally have Re_u: "Re u \ x/2" using elim by simp have "norm (stirling_integral m u / (u - of_real x) ^ Suc j) \ c / Re u ^ m / (x / 2) ^ Suc j" using Re_u elim unfolding norm_divide norm_power dist by (intro divide_right_mono zero_le_power c) simp_all also have "\ \ c / (x/2) ^ m / (x / 2) ^ Suc j" using c_nonneg elim Re_u by (intro divide_right_mono divide_left_mono power_mono) simp_all also have "\ = B / x ^ (m + Suc j)" using elim by (simp add: B_def field_simps power_add) finally show "norm (stirling_integral m u / (u - of_real x) ^ Suc j) \ B / x ^ (m + Suc j)" . qed (insert elim c_nonneg, auto simp: B_def simp del: power_Suc) hence "cmod ((deriv ^^ j) (stirling_integral m) (of_real x)) \ B' / x ^ (j + m)" using elim by (simp add: field_simps norm_divide norm_mult norm_power B'_def) with elim m show ?case by (simp_all add: add_ac deriv_stirling_integral_complex_of_real) qed thus ?thesis by (rule bigoI) qed definition stirling_sum where "stirling_sum j m x = (-1) ^ j * (\k = 1..k\m. (of_real (bernoulli' k) * pochhammer (of_nat (Suc k)) (j - 1) * inverse x ^ (k + j)))" lemma stirling_sum_complex_of_real: "stirling_sum j m (complex_of_real x) = complex_of_real (stirling_sum j m x)" by (simp add: stirling_sum_def pochhammer_of_real [symmetric] del: of_nat_Suc) lemma stirling_sum'_complex_of_real: "stirling_sum' j m (complex_of_real x) = complex_of_real (stirling_sum' j m x)" by (simp add: stirling_sum'_def pochhammer_of_real [symmetric] del: of_nat_Suc) lemma has_field_derivative_stirling_sum_complex [derivative_intros]: "Re x > 0 \ (stirling_sum j m has_field_derivative stirling_sum (Suc j) m x) (at x)" unfolding stirling_sum_def [abs_def] sum_distrib_left by (rule DERIV_sum) (auto intro!: derivative_eq_intros simp del: of_nat_Suc simp: pochhammer_Suc power_diff) lemma has_field_derivative_stirling_sum_real [derivative_intros]: "x > (0::real) \ (stirling_sum j m has_field_derivative stirling_sum (Suc j) m x) (at x)" unfolding stirling_sum_def [abs_def] sum_distrib_left by (rule DERIV_sum) (auto intro!: derivative_eq_intros simp del: of_nat_Suc simp: pochhammer_Suc power_diff) lemma has_field_derivative_stirling_sum'_complex [derivative_intros]: assumes "j > 0" "Re x > 0" shows "(stirling_sum' j m has_field_derivative stirling_sum' (Suc j) m x) (at x)" proof (cases j) case (Suc j') from assms have [simp]: "x \ 0" by auto define c where "c = (\n. (-1) ^ Suc j * complex_of_real (bernoulli' n) * pochhammer (of_nat (Suc n)) j')" define T where "T = (\n x. c n * inverse x ^ (j + n))" define T' where "T' = (\n x. - (of_nat (j + n)) * c n * inverse x ^ (Suc (j + n)))" have "((\x. \k\m. T k x) has_field_derivative (\k\m. T' k x)) (at x)" using assms Suc by (intro DERIV_sum) (auto simp: T_def T'_def intro!: derivative_eq_intros simp: field_simps power_add [symmetric] simp del: of_nat_Suc power_Suc of_nat_add) also have "(\x. (\k\m. T k x)) = stirling_sum' j m" by (simp add: Suc T_def c_def stirling_sum'_def fun_eq_iff add_ac mult.assoc sum_distrib_left) also have "(\k\m. T' k x) = stirling_sum' (Suc j) m x" by (simp add: T'_def c_def Suc stirling_sum'_def sum_distrib_left sum_distrib_right algebra_simps pochhammer_Suc) finally show ?thesis . qed (insert assms, simp_all) lemma has_field_derivative_stirling_sum'_real [derivative_intros]: assumes "j > 0" "x > (0::real)" shows "(stirling_sum' j m has_field_derivative stirling_sum' (Suc j) m x) (at x)" proof (cases j) case (Suc j') from assms have [simp]: "x \ 0" by auto define c where "c = (\n. (-1) ^ Suc j * (bernoulli' n) * pochhammer (of_nat (Suc n)) j')" define T where "T = (\n x. c n * inverse x ^ (j + n))" define T' where "T' = (\n x. - (of_nat (j + n)) * c n * inverse x ^ (Suc (j + n)))" have "((\x. \k\m. T k x) has_field_derivative (\k\m. T' k x)) (at x)" using assms Suc by (intro DERIV_sum) (auto simp: T_def T'_def intro!: derivative_eq_intros simp: field_simps power_add [symmetric] simp del: of_nat_Suc power_Suc of_nat_add) also have "(\x. (\k\m. T k x)) = stirling_sum' j m" by (simp add: Suc T_def c_def stirling_sum'_def fun_eq_iff add_ac mult.assoc sum_distrib_left) also have "(\k\m. T' k x) = stirling_sum' (Suc j) m x" by (simp add: T'_def c_def Suc stirling_sum'_def sum_distrib_left sum_distrib_right algebra_simps pochhammer_Suc) finally show ?thesis . qed (insert assms, simp_all) lemma higher_deriv_stirling_sum_complex: "Re x > 0 \ (deriv ^^ i) (stirling_sum j m) x = stirling_sum (i + j) m x" proof (induction i arbitrary: x) case (Suc i) have "deriv ((deriv ^^ i) (stirling_sum j m)) x = deriv (stirling_sum (i + j) m) x" using eventually_nhds_in_open[of "{x. Re x > 0}" x] Suc.prems by (intro deriv_cong_ev refl) (auto elim!: eventually_mono simp: open_halfspace_Re_gt Suc.IH) also from Suc.prems have "\ = stirling_sum (Suc (i + j)) m x" by (intro DERIV_imp_deriv has_field_derivative_stirling_sum_complex) finally show ?case by simp qed simp_all definition Polygamma_approx :: "nat \ nat \ 'a \ 'a :: {real_normed_field, ln}" where "Polygamma_approx j m = (deriv ^^ j) (\x. (x - 1 / 2) * ln x - x + of_real (ln (2 * pi)) / 2 + stirling_sum 0 m x)" lemma Polygamma_approx_Suc: "Polygamma_approx (Suc j) m = deriv (Polygamma_approx j m)" by (simp add: Polygamma_approx_def) lemma Polygamma_approx_0: "Polygamma_approx 0 m x = (x - 1/2) * ln x - x + of_real (ln (2*pi)) / 2 + stirling_sum 0 m x" by (simp add: Polygamma_approx_def) lemma Polygamma_approx_1_complex: "Re x > 0 \ Polygamma_approx (Suc 0) m x = ln x - 1 / (2*x) + stirling_sum (Suc 0) m x" unfolding Polygamma_approx_Suc Polygamma_approx_0 by (intro DERIV_imp_deriv) (auto intro!: derivative_eq_intros elim!: nonpos_Reals_cases simp: field_simps) lemma Polygamma_approx_1_real: "x > (0 :: real) \ Polygamma_approx (Suc 0) m x = ln x - 1 / (2*x) + stirling_sum (Suc 0) m x" unfolding Polygamma_approx_Suc Polygamma_approx_0 by (intro DERIV_imp_deriv) (auto intro!: derivative_eq_intros elim!: nonpos_Reals_cases simp: field_simps) lemma stirling_sum_2_conv_stirling_sum'_1: fixes x :: "'a :: {real_div_algebra, field_char_0}" assumes "m > 0" "x \ 0" shows "stirling_sum' 1 m x = 1 / x + 1 / (2 * x^2) + stirling_sum 2 m x" proof - have pochhammer_2: "pochhammer (of_nat k) 2 = of_nat k * of_nat (Suc k)" for k by (simp add: pochhammer_Suc eval_nat_numeral add_ac) have "stirling_sum 2 m x = (\k = Suc 0.. = (\k=0.. = (\k=0.. = (\k\m. of_real (bernoulli' k) * inverse x ^ Suc k)" by (intro sum.cong) auto also have "\ = stirling_sum' 1 m x" by (simp add: stirling_sum'_def) finally show ?thesis by (simp add: add_ac) qed lemma Polygamma_approx_2_real: assumes "x > (0::real)" "m > 0" shows "Polygamma_approx (Suc (Suc 0)) m x = stirling_sum' 1 m x" proof - have "Polygamma_approx (Suc (Suc 0)) m x = deriv (Polygamma_approx (Suc 0) m) x" by (simp add: Polygamma_approx_Suc) also have "\ = deriv (\x. ln x - 1 / (2*x) + stirling_sum (Suc 0) m x) x" using eventually_nhds_in_open[of "{0<..}" x] assms by (intro deriv_cong_ev) (auto elim!: eventually_mono simp: Polygamma_approx_1_real) also have "\ = 1 / x + 1 / (2*x^2) + stirling_sum (Suc (Suc 0)) m x" using assms by (intro DERIV_imp_deriv) (auto intro!: derivative_eq_intros elim!: nonpos_Reals_cases simp: field_simps power2_eq_square) also have "\ = stirling_sum' 1 m x" using stirling_sum_2_conv_stirling_sum'_1[of m x] assms by (simp add: eval_nat_numeral) finally show ?thesis . qed lemma Polygamma_approx_2_complex: assumes "Re x > 0" "m > 0" shows "Polygamma_approx (Suc (Suc 0)) m x = stirling_sum' 1 m x" proof - have "Polygamma_approx (Suc (Suc 0)) m x = deriv (Polygamma_approx (Suc 0) m) x" by (simp add: Polygamma_approx_Suc) also have "\ = deriv (\x. ln x - 1 / (2*x) + stirling_sum (Suc 0) m x) x" using eventually_nhds_in_open[of "{s. Re s > 0}" x] assms by (intro deriv_cong_ev) (auto simp: open_halfspace_Re_gt elim!: eventually_mono simp: Polygamma_approx_1_complex) also have "\ = 1 / x + 1 / (2*x^2) + stirling_sum (Suc (Suc 0)) m x" using assms by (intro DERIV_imp_deriv) (auto intro!: derivative_eq_intros elim!: nonpos_Reals_cases simp: field_simps power2_eq_square) also have "\ = stirling_sum' 1 m x" using stirling_sum_2_conv_stirling_sum'_1[of m x] assms by (subst stirling_sum_2_conv_stirling_sum'_1) (auto simp: eval_nat_numeral) finally show ?thesis . qed lemma Polygamma_approx_ge_2_real: assumes "x > (0::real)" "m > 0" shows "Polygamma_approx (Suc (Suc j)) m x = stirling_sum' (Suc j) m x" using assms(1) proof (induction j arbitrary: x) case (0 x) with assms show ?case by (simp add: Polygamma_approx_2_real) next case (Suc j x) have "Polygamma_approx (Suc (Suc (Suc j))) m x = deriv (Polygamma_approx (Suc (Suc j)) m) x" by (simp add: Polygamma_approx_Suc) also have "\ = deriv (stirling_sum' (Suc j) m) x" using eventually_nhds_in_open[of "{0<..}" x] Suc.prems by (intro deriv_cong_ev refl) (auto elim!: eventually_mono simp: Suc.IH) also have "\ = stirling_sum' (Suc (Suc j)) m x" using Suc.prems by (intro DERIV_imp_deriv derivative_intros) simp_all finally show ?case . qed lemma Polygamma_approx_ge_2_complex: assumes "Re x > 0" "m > 0" shows "Polygamma_approx (Suc (Suc j)) m x = stirling_sum' (Suc j) m x" using assms(1) proof (induction j arbitrary: x) case (0 x) with assms show ?case by (simp add: Polygamma_approx_2_complex) next case (Suc j x) have "Polygamma_approx (Suc (Suc (Suc j))) m x = deriv (Polygamma_approx (Suc (Suc j)) m) x" by (simp add: Polygamma_approx_Suc) also have "\ = deriv (stirling_sum' (Suc j) m) x" using eventually_nhds_in_open[of "{x. Re x > 0}" x] Suc.prems by (intro deriv_cong_ev refl) (auto elim!: eventually_mono simp: Suc.IH open_halfspace_Re_gt) also have "\ = stirling_sum' (Suc (Suc j)) m x" using Suc.prems by (intro DERIV_imp_deriv derivative_intros) simp_all finally show ?case . qed lemma Polygamma_approx_complex_of_real: assumes "x > 0" "m > 0" shows "Polygamma_approx j m (complex_of_real x) = of_real (Polygamma_approx j m x)" proof (cases j) case 0 with assms show ?thesis by (simp add: Polygamma_approx_0 Ln_of_real stirling_sum_complex_of_real) next case [simp]: (Suc j') thus ?thesis proof (cases j') case 0 with assms show ?thesis by (simp add: Polygamma_approx_1_complex Polygamma_approx_1_real stirling_sum_complex_of_real Ln_of_real) next case (Suc j'') with assms show ?thesis by (simp add: Polygamma_approx_ge_2_complex Polygamma_approx_ge_2_real stirling_sum'_complex_of_real) qed qed lemma higher_deriv_Polygamma_approx [simp]: "(deriv ^^ j) (Polygamma_approx i m) = Polygamma_approx (j + i) m" by (simp add: Polygamma_approx_def funpow_add) lemma stirling_sum_holomorphic [holomorphic_intros]: "0 \ A \ stirling_sum j m holomorphic_on A" unfolding stirling_sum_def by (intro holomorphic_intros) auto lemma Polygamma_approx_holomorphic [holomorphic_intros]: "Polygamma_approx j m holomorphic_on {s. Re s > 0}" unfolding Polygamma_approx_def by (intro holomorphic_intros) (auto simp: open_halfspace_Re_gt elim!: nonpos_Reals_cases) lemma higher_deriv_lnGamma_stirling: assumes m: "m > 0" shows "(\x::real. (deriv ^^ j) ln_Gamma x - Polygamma_approx j m x) \ O(\x. 1 / x ^ (m + j))" proof - have "eventually (\x. \(deriv ^^ j) ln_Gamma x - Polygamma_approx j m x\ = inverse (real m) * \(deriv ^^ j) (stirling_integral m) x\) at_top" using eventually_gt_at_top[of "0::real"] proof eventually_elim case (elim x) note x = this have "(deriv ^^ j) (\x. ln_Gamma x - Polygamma_approx 0 m x) (complex_of_real x) = (deriv ^^ j) (\x. (-inverse (of_nat m)) * stirling_integral m x) (complex_of_real x)" using eventually_nhds_in_open[of "{s. Re s > 0}" x] x m by (intro higher_deriv_cong_ev refl) (auto elim!: eventually_mono simp: ln_Gamma_stirling_complex Polygamma_approx_def field_simps open_halfspace_Re_gt stirling_sum_def) also have "\ = - inverse (of_nat m) * (deriv ^^ j) (stirling_integral m) (of_real x)" using x m by (intro higher_deriv_cmult[of _ "{s. Re s > 0}"] stirling_integral_holomorphic) (auto simp: open_halfspace_Re_gt) also have "(deriv ^^ j) (\x. ln_Gamma x - Polygamma_approx 0 m x) (complex_of_real x) = (deriv ^^ j) ln_Gamma (of_real x) - (deriv ^^ j) (Polygamma_approx 0 m) (of_real x)" using x by (intro higher_deriv_diff[of _ "{s. Re s > 0}"]) (auto intro!: holomorphic_intros elim!: nonpos_Reals_cases simp: open_halfspace_Re_gt) also have "(deriv ^^ j) (Polygamma_approx 0 m) (complex_of_real x) = of_real (Polygamma_approx j m x)" using x m by (simp add: Polygamma_approx_complex_of_real) also have "norm (- inverse (of_nat m) * (deriv ^^ j) (stirling_integral m) (complex_of_real x)) = inverse (real m) * \(deriv ^^ j) (stirling_integral m) x\" using x m by (simp add: norm_mult norm_inverse deriv_stirling_integral_complex_of_real) also have "(deriv ^^ j) ln_Gamma (complex_of_real x) = of_real ((deriv ^^ j) ln_Gamma x)" using x by (simp add: higher_deriv_ln_Gamma_complex_of_real) also have "norm (\ - of_real (Polygamma_approx j m x)) = \(deriv ^^ j) ln_Gamma x - Polygamma_approx j m x\" by (simp only: of_real_diff [symmetric] norm_of_real) finally show ?case . qed from bigthetaI_cong[OF this] m have "(\x::real. (deriv ^^ j) ln_Gamma x - Polygamma_approx j m x) \ \(\x. (deriv ^^ j) (stirling_integral m) x)" by simp also have "(\x::real. (deriv ^^ j) (stirling_integral m) x) \ O(\x. 1 / x ^ (m + j))" using m by (rule deriv_stirling_integral_real_bound) finally show ?thesis . qed lemma Polygamma_approx_1_real': assumes x: "(x::real) > 0" and m: "m > 0" shows "Polygamma_approx 1 m x = ln x - (\k = Suc 0..m. bernoulli' k * inverse x ^ k / real k)" proof - have "Polygamma_approx 1 m x = ln x - (1 / (2 * x) + (\k=Suc 0..k=Suc 0.. = (\k=0.. = (\k = Suc 0..m. bernoulli' k * inverse x ^ k / real k)" using assms by (subst sum.shift_bounds_Suc_ivl [symmetric]) (simp add: atLeastLessThanSuc_atLeastAtMost) finally show ?thesis . qed theorem assumes m: "m > 0" shows ln_Gamma_real_asymptotics: "(\x. ln_Gamma x - ((x - 1 / 2) * ln x - x + ln (2 * pi) / 2 + (\k = 1.. O(\x. 1 / x ^ m)" (is ?th1) and Digamma_real_asymptotics: "(\x. Digamma x - (ln x - (\k=1..m. bernoulli' k / real k / x ^ k))) \ O(\x. 1 / (x ^ Suc m))" (is ?th2) and Polygamma_real_asymptotics: "j > 0 \ (\x. Polygamma j x - (- 1) ^ Suc j * (\k\m. bernoulli' k * pochhammer (real (Suc k)) (j - 1) / x ^ (k + j))) \ O(\x. 1 / x ^ (m+j+1))" (is "_ \ ?th3") proof - define G :: "nat \ real \ real" where "G = (\m. if m = 0 then ln_Gamma else Polygamma (m - 1))" have *: "(\x. G j x - h x) \ O(\x. 1 / x ^ (m + j))" if "\x::real. x > 0 \ Polygamma_approx j m x = h x" for j h proof - have "(\x. G j x - h x) \ \(\x. (deriv ^^ j) ln_Gamma x - Polygamma_approx j m x)" (is "_ \ \(?f)") using that by (intro bigthetaI_cong) (auto intro: eventually_mono[OF eventually_gt_at_top[of "0::real"]] simp del: funpow.simps simp: higher_deriv_ln_Gamma_real G_def) also have "?f \ O(\x::real. 1 / x ^ (m + j))" using m by (rule higher_deriv_lnGamma_stirling) finally show ?thesis . qed note [[simproc del: simplify_landau_sum]] from *[OF Polygamma_approx_0] assms show ?th1 by (simp add: G_def Polygamma_approx_0 stirling_sum_def field_simps) from *[OF Polygamma_approx_1_real'] assms show ?th2 by (simp add: G_def field_simps) assume j: "j > 0" from *[OF Polygamma_approx_ge_2_real, of "j - 1"] assms j show ?th3 by (simp add: G_def stirling_sum'_def power_add power_diff field_simps) qed end diff --git a/thys/Stirling_Formula/ROOT b/thys/Stirling_Formula/ROOT --- a/thys/Stirling_Formula/ROOT +++ b/thys/Stirling_Formula/ROOT @@ -1,13 +1,13 @@ chapter AFP -session Stirling_Formula (AFP) = "HOL-Analysis" + +session Stirling_Formula (AFP) = Bernoulli + options [timeout = 600] sessions Bernoulli Landau_Symbols theories Stirling_Formula Ln_Gamma_Asymptotics document_files "root.tex" "root.bib" diff --git a/thys/Transcendence_Series_Hancl_Rucki/ROOT b/thys/Transcendence_Series_Hancl_Rucki/ROOT --- a/thys/Transcendence_Series_Hancl_Rucki/ROOT +++ b/thys/Transcendence_Series_Hancl_Rucki/ROOT @@ -1,12 +1,12 @@ chapter AFP -session Transcendence_Series_Hancl_Rucki (AFP) = "HOL-Analysis" + +session Transcendence_Series_Hancl_Rucki (AFP) = "HOL-Complex_Analysis" + options [timeout = 1200] sessions Prime_Number_Theorem theories Transcendence_Series document_files "root.tex" "root.bib" diff --git a/thys/Winding_Number_Eval/Cauchy_Index_Theorem.thy b/thys/Winding_Number_Eval/Cauchy_Index_Theorem.thy --- a/thys/Winding_Number_Eval/Cauchy_Index_Theorem.thy +++ b/thys/Winding_Number_Eval/Cauchy_Index_Theorem.thy @@ -1,4682 +1,4682 @@ (* Author: Wenda Li *) section \Cauchy's index theorem\ theory Cauchy_Index_Theorem imports - "HOL-Analysis.Multivariate_Analysis" + "HOL-Complex_Analysis.Complex_Analysis" "Sturm_Tarski.Sturm_Tarski" "HOL-Computational_Algebra.Fundamental_Theorem_Algebra" Missing_Transcendental Missing_Algebraic Missing_Analysis begin text \This theory formalises Cauchy indices on the complex plane and relate them to winding numbers\ subsection \Misc\ (*refined version of the library one with the same name by dropping unnecessary assumptions*) lemma atMostAtLeast_subset_convex: fixes C :: "real set" assumes "convex C" and "x \ C" "y \ C" shows "{x .. y} \ C" proof safe fix z assume z: "z \ {x .. y}" have "z \ C" if *: "x < z" "z < y" proof - let ?\ = "(y - z) / (y - x)" have "0 \ ?\" "?\ \ 1" using assms * by (auto simp: field_simps) then have comb: "?\ * x + (1 - ?\) * y \ C" using assms iffD1[OF convex_alt, rule_format, of C y x ?\] by (simp add: algebra_simps) have "?\ * x + (1 - ?\) * y = (y - z) * x / (y - x) + (1 - (y - z) / (y - x)) * y" by (auto simp: field_simps) also have "\ = ((y - z) * x + (y - x - (y - z)) * y) / (y - x)" using * by (simp only: add_divide_distrib) (auto simp: field_simps) also have "\ = z" using assms * by (auto simp: field_simps) finally show ?thesis using comb by auto qed then show "z \ C" using z assms by (auto simp: le_less) qed lemma arg_elim: "f x \ x= y \ f y" by auto lemma arg_elim2: "f x1 x2 \ x1= y1 \x2=y2 \ f y1 y2" by auto lemma arg_elim3: "\f x1 x2 x3;x1= y1;x2=y2;x3=y3 \ \ f y1 y2 y3" by auto lemma IVT_strict: fixes f :: "'a::linear_continuum_topology \ 'b::linorder_topology" assumes "(f a > y \ y > f b) \ (f a < y \ y < f b)" "ax. a < x \ x < b \ f x = y" by (metis IVT' IVT2' assms(1) assms(2) assms(3) linorder_neq_iff order_le_less order_less_imp_le) lemma (in dense_linorder) atLeastAtMost_subseteq_greaterThanLessThan_iff: "{a .. b} \ { c <..< d } \ (a \ b \ c < a \ b < d)" using dense[of "a" "min c b"] dense[of "max a d" "b"] by (force simp: subset_eq Ball_def not_less[symmetric]) lemma Re_winding_number_half_right: assumes "\p\path_image \. Re p \ Re z" and "valid_path \" and "z\path_image \" shows "Re(winding_number \ z) = (Im (Ln (pathfinish \ - z)) - Im (Ln (pathstart \ - z)))/(2*pi)" proof - define g where "g=(\t. \ t - z)" define st fi where "st\pathstart g" and "fi\pathfinish g" have "valid_path g" "0\path_image g" and pos_img:"\p\path_image g. Re p \ 0" unfolding g_def subgoal using assms(2) by auto subgoal using assms(3) by auto subgoal using assms(1) by fastforce done have "(inverse has_contour_integral Ln fi - Ln st) g" unfolding fi_def st_def proof (rule contour_integral_primitive[OF _ \valid_path g\,of " - \\<^sub>\\<^sub>0"]) fix x::complex assume "x \ - \\<^sub>\\<^sub>0" then have "(Ln has_field_derivative inverse x) (at x)" using has_field_derivative_Ln by auto then show "(Ln has_field_derivative inverse x) (at x within - \\<^sub>\\<^sub>0)" using has_field_derivative_at_within by auto next show "path_image g \ - \\<^sub>\\<^sub>0" using pos_img \0\path_image g\ by (metis ComplI antisym assms(3) complex_nonpos_Reals_iff complex_surj subsetI zero_complex.code) qed then have winding_eq:"2*pi*\*winding_number g 0 = (Ln fi - Ln st)" using has_contour_integral_winding_number[OF \valid_path g\ \0\path_image g\ ,simplified,folded inverse_eq_divide] has_contour_integral_unique by auto have "Re(winding_number g 0) = (Im (Ln fi) - Im (Ln st))/(2*pi)" (is "?L=?R") proof - have "?L = Re((Ln fi - Ln st)/(2*pi*\))" using winding_eq[symmetric] by auto also have "... = ?R" by (metis Im_divide_of_real Im_i_times complex_i_not_zero minus_complex.simps(2) mult.commute mult_divide_mult_cancel_left_if times_divide_eq_right) finally show ?thesis . qed then show ?thesis unfolding g_def fi_def st_def using winding_number_offset by simp qed lemma Re_winding_number_half_upper: assumes pimage:"\p\path_image \. Im p \ Im z" and "valid_path \" and "z\path_image \" shows "Re(winding_number \ z) = (Im (Ln (\*z - \*pathfinish \)) - Im (Ln (\*z - \*pathstart \ )))/(2*pi)" proof - define \' where "\'=(\t. - \ * (\ t - z) + z)" have "Re (winding_number \' z) = (Im (Ln (pathfinish \' - z)) - Im (Ln (pathstart \' - z))) / (2 * pi)" unfolding \'_def apply (rule Re_winding_number_half_right) subgoal using pimage unfolding path_image_def by auto subgoal apply (rule valid_path_compose_holomorphic[OF \valid_path \\,of "\x. -\ * (x-z) + z" UNIV , unfolded comp_def]) by (auto intro!:holomorphic_intros) subgoal using \z\path_image \\ unfolding path_image_def by auto done moreover have "winding_number \' z = winding_number \ z" proof - define f where "f=(\x. -\ * (x-z) + z)" define c where "c= 1 / (complex_of_real (2 * pi) * \)" have "winding_number \' z = winding_number (f o \) z" unfolding \'_def comp_def f_def by auto also have "... = c * contour_integral \ (\w. deriv f w / (f w - z))" unfolding c_def proof (rule winding_number_comp[of UNIV]) show "z \ path_image (f \ \)" using \z\path_image \\ unfolding f_def path_image_def by auto qed (auto simp add:f_def \valid_path \\ intro!:holomorphic_intros) also have "... = c * contour_integral \ (\w. 1 / (w - z))" proof - have "deriv f x = -\" for x unfolding f_def by (auto intro!:derivative_eq_intros DERIV_imp_deriv) then show ?thesis unfolding f_def c_def by (auto simp add:field_simps divide_simps intro!:arg_cong2[where f=contour_integral]) qed also have "... = winding_number \ z" using winding_number_valid_path[OF \valid_path \\ \z\path_image \\,folded c_def] by simp finally show ?thesis . qed moreover have "pathfinish \' = z+ \*z -\* pathfinish \" "pathstart \' = z+ \*z -\*pathstart \" unfolding \'_def path_defs by (auto simp add:algebra_simps) ultimately show ?thesis by auto qed lemma Re_winding_number_half_lower: assumes pimage:"\p\path_image \. Im p \ Im z" and "valid_path \" and "z\path_image \" shows "Re(winding_number \ z) = (Im (Ln ( \*pathfinish \ - \*z)) - Im (Ln (\*pathstart \ - \*z)))/(2*pi)" proof - define \' where "\'=(\t. \ * (\ t - z) + z)" have "Re (winding_number \' z) = (Im (Ln (pathfinish \' - z)) - Im (Ln (pathstart \' - z))) / (2 * pi)" unfolding \'_def apply (rule Re_winding_number_half_right) subgoal using pimage unfolding path_image_def by auto subgoal apply (rule valid_path_compose_holomorphic[OF \valid_path \\,of "\x. \ * (x-z) + z" UNIV , unfolded comp_def]) by (auto intro!:holomorphic_intros) subgoal using \z\path_image \\ unfolding path_image_def by auto done moreover have "winding_number \' z = winding_number \ z" proof - define f where "f=(\x. \ * (x-z) + z)" define c where "c= 1 / (complex_of_real (2 * pi) * \)" have "winding_number \' z = winding_number (f o \) z" unfolding \'_def comp_def f_def by auto also have "... = c * contour_integral \ (\w. deriv f w / (f w - z))" unfolding c_def proof (rule winding_number_comp[of UNIV]) show "z \ path_image (f \ \)" using \z\path_image \\ unfolding f_def path_image_def by auto qed (auto simp add:f_def \valid_path \\ intro!:holomorphic_intros) also have "... = c * contour_integral \ (\w. 1 / (w - z))" proof - have "deriv f x = \" for x unfolding f_def by (auto intro!:derivative_eq_intros DERIV_imp_deriv) then show ?thesis unfolding f_def c_def by (auto simp add:field_simps divide_simps intro!:arg_cong2[where f=contour_integral]) qed also have "... = winding_number \ z" using winding_number_valid_path[OF \valid_path \\ \z\path_image \\,folded c_def] by simp finally show ?thesis . qed moreover have "pathfinish \' = z+ \* pathfinish \ - \*z" "pathstart \' = z+ \*pathstart \ - \*z" unfolding \'_def path_defs by (auto simp add:algebra_simps) ultimately show ?thesis by auto qed lemma Re_winding_number_half_left: assumes neg_img:"\p\path_image \. Re p \ Re z" and "valid_path \" and "z\path_image \" shows "Re(winding_number \ z) = (Im (Ln (z - pathfinish \)) - Im (Ln (z - pathstart \ )))/(2*pi)" proof - define \' where "\'\(\t. 2*z - \ t)" have "Re (winding_number \' z) = (Im (Ln (pathfinish \' - z)) - Im (Ln (pathstart \' - z))) / (2 * pi)" unfolding \'_def apply (rule Re_winding_number_half_right) subgoal using neg_img unfolding path_image_def by auto subgoal apply (rule valid_path_compose_holomorphic[OF \valid_path \\,of "\t. 2*z-t" UNIV, unfolded comp_def]) by (auto intro:holomorphic_intros) subgoal using \z\path_image \\ unfolding path_image_def by auto done moreover have "winding_number \' z = winding_number \ z" proof - define f where "f=(\t. 2*z-t)" define c where "c= 1 / (complex_of_real (2 * pi) * \)" have "winding_number \' z = winding_number (f o \) z" unfolding \'_def comp_def f_def by auto also have "... = c * contour_integral \ (\w. deriv f w / (f w - z))" unfolding c_def proof (rule winding_number_comp[of UNIV]) show "z \ path_image (f \ \)" using \z\path_image \\ unfolding f_def path_image_def by auto qed (auto simp add:f_def \valid_path \\ intro:holomorphic_intros) also have "... = c * contour_integral \ (\w. 1 / (w - z))" unfolding f_def c_def by (auto simp add:field_simps divide_simps intro!:arg_cong2[where f=contour_integral]) also have "... = winding_number \ z" using winding_number_valid_path[OF \valid_path \\ \z\path_image \\,folded c_def] by simp finally show ?thesis . qed moreover have "pathfinish \' = 2*z - pathfinish \" "pathstart \' = 2*z - pathstart \" unfolding \'_def path_defs by auto ultimately show ?thesis by auto qed lemma continuous_on_open_Collect_neq: fixes f g :: "'a::topological_space \ 'b::t2_space" assumes f: "continuous_on S f" and g: "continuous_on S g" and "open S" shows "open {x\S. f x \ g x}" proof (rule topological_space_class.openI) fix t assume "t \ {x\S. f x \ g x}" then obtain U0 V0 where "open U0" "open V0" "f t \ U0" "g t \ V0" "U0 \ V0 = {}" "t\S" by (auto simp add: separation_t2) obtain U1 where "open U1" "t \ U1" "\y\(S \ U1). f y \ U0" using f[unfolded continuous_on_topological,rule_format,OF \t\S\ \open U0\ \f t \U0\] by auto obtain V1 where "open V1" "t \ V1" "\y\(S \ V1). g y \ V0" using g[unfolded continuous_on_topological,rule_format,OF \t\S\ \open V0\ \g t \V0\] by auto define T where "T=V1 \ U1 \ S" have "open T" unfolding T_def using \open U1\ \open V1\ \open S\ by auto moreover have "t \ T" unfolding T_def using \t\U1\ \t\V1\ \t\S\ by auto moreover have "T \ {x \ S. f x \ g x}" unfolding T_def using \U0 \ V0 = {}\ \\y\S \ U1. f y \ U0\ \\y\S \ V1. g y \ V0\ by auto ultimately show "\T. open T \ t \ T \ T \ {x \ S. f x \ g x}" by auto qed subsection \Sign at a filter\ (*Relaxation in types could be done in the future.*) definition has_sgnx::"(real \ real) \ real \ real filter \ bool" (infixr "has'_sgnx" 55) where "(f has_sgnx c) F= (eventually (\x. sgn(f x) = c) F)" definition sgnx_able (infixr "sgnx'_able" 55) where "(f sgnx_able F) = (\c. (f has_sgnx c) F)" definition sgnx where "sgnx f F = (SOME c. (f has_sgnx c) F)" lemma has_sgnx_eq_rhs: "(f has_sgnx x) F \ x = y \ (f has_sgnx y) F" by simp named_theorems sgnx_intros "introduction rules for has_sgnx" setup \ Global_Theory.add_thms_dynamic (@{binding sgnx_eq_intros}, fn context => Named_Theorems.get (Context.proof_of context) @{named_theorems sgnx_intros} |> map_filter (try (fn thm => @{thm has_sgnx_eq_rhs} OF [thm]))) \ lemma sgnx_able_sgnx:"f sgnx_able F \ (f has_sgnx (sgnx f F)) F" unfolding sgnx_able_def sgnx_def using someI_ex by metis lemma has_sgnx_imp_sgnx_able[elim]: "(f has_sgnx c) F \ f sgnx_able F" unfolding sgnx_able_def by auto lemma has_sgnx_unique: assumes "F\bot" "(f has_sgnx c1) F" "(f has_sgnx c2) F" shows "c1=c2" proof (rule ccontr) assume "c1 \ c2 " have "eventually (\x. sgn(f x) = c1 \ sgn(f x) = c2) F" using assms unfolding has_sgnx_def eventually_conj_iff by simp then have "eventually (\_. c1 = c2) F" by (elim eventually_mono,auto) then have "eventually (\_. False) F" using \c1 \ c2\ by auto then show False using \F \ bot\ eventually_False by auto qed lemma has_sgnx_imp_sgnx[elim]: "(f has_sgnx c) F \F\bot \ sgnx f F = c" using has_sgnx_unique sgnx_def by auto lemma has_sgnx_const[simp,sgnx_intros]: "((\_. c) has_sgnx sgn c) F" by (simp add: has_sgnx_def) lemma finite_sgnx_at_left_at_right: assumes "finite {t. f t=0 \ a t{a<..0" "f sgnx_able (at_right x)" "sgnx f (at_right x)\0" proof - define ls where "ls \ {t. (f t=0 \ t\s) \ at(if ls = {} then (a+x)/2 else (Max ls + x)/2)" have "finite ls" proof - have "{t. f t=0 \ a t {t. f t=0 \ a t a ts \ a t a t {t. t\s \ a t a l a l{}" proof - have "Max ls \ ls" using assms(1,3) that \finite ls\ apply (intro linorder_class.Max_in) by auto then have "a Max ls < x" unfolding ls_def by auto then show ?thesis unfolding l_def using that x by auto qed ultimately show "l0" when t:"t\{l..a" using t \l>a\ by (meson atLeastLessThan_iff less_le_trans) then have "t\ls" using that t unfolding ls_def by auto then show False using True by auto qed then show ?thesis by auto next case False have "t>Max ls" using that False \l unfolding l_def by auto have False when "f t=0" proof - have "t>a" using t \l>a\ by (meson atLeastLessThan_iff less_le_trans) then have "t\ls" using that t unfolding ls_def by auto then have "t\Max ls" using \finite ls\ by auto then show False using \t>Max ls\ by auto qed then show ?thesis by auto qed have "(f has_sgnx sgn (f l)) (at_left x)" unfolding has_sgnx_def proof (rule eventually_at_leftI[OF _ \l]) fix t assume t:"t\{l<..a" "tl>a\ x by (meson greaterThanLessThan_iff less_trans)+ have False when "f t = 0" using noroot t that by auto moreover have False when "f l=0" using noroot t that by auto moreover have False when "f l>0 \ f t<0 \ f l <0 \ f t >0" proof - have False when "{l..t} \ s \{}" proof - obtain t' where t':"t'\{l..t}" "t'\s" using \{l..t} \ s \ {}\ by blast then have "a t'a < l\ atLeastAtMost_iff greaterThanLessThan_iff le_less less_trans t) then have "t'\ls" unfolding ls_def using \t'\s\ by auto then have "t'\Max ls" using \finite ls\ by auto moreover have "Max lsl \t'\ls\ \finite ls\ unfolding l_def by (auto simp add:ls_def ) ultimately show False using t'(1) by auto qed moreover have "{l..t} \ {a<..x>l. x < t \ f x = 0" apply (intro IVT_strict) using that t assms(2) by auto then obtain t' where "l{l..f t'=0\ by auto qed ultimately show "sgn (f t) = sgn (f l)" by (metis le_less not_less sgn_if) qed then show "f sgnx_able (at_left x)" by auto show "sgnx f (at_left x)\0" using noroot[of l,simplified] \(f has_sgnx sgn (f l)) (at_left x)\ by (simp add: has_sgnx_imp_sgnx sgn_if) next define rs where "rs \ {t. (f t=0 \ t\s) \ x t(if rs = {} then (x+b)/2 else (Min rs + x)/2)" have "finite rs" proof - have "{t. f t=0 \ x t {t. f t=0 \ a t x ts \ x t x t {t. t\s \ x tx" "ax \ a rx \ a r{}" proof - have "Min rs \ rs" using assms(1,3) that \finite rs\ apply (intro linorder_class.Min_in) by auto then have "x Min rs < b" unfolding rs_def by auto then show ?thesis unfolding r_def using that x by auto qed ultimately show "r>x" "a0" when t:"t\{x<..r}" for t proof (cases "rs = {}") case True have False when "f t=0" proof - have "tr using greaterThanAtMost_iff by fastforce then have "t\rs" using that t unfolding rs_def by auto then show False using True by auto qed then show ?thesis by auto next case False have "tr>x\ unfolding r_def by auto have False when "f t=0" proof - have "tr by (metis greaterThanAtMost_iff le_less less_trans) then have "t\rs" using that t unfolding rs_def by auto then have "t\Min rs" using \finite rs\ by auto then show False using \t by auto qed then show ?thesis by auto qed have "(f has_sgnx sgn (f r)) (at_right x)" unfolding has_sgnx_def proof (rule eventually_at_rightI[OF _ \r>x\]) fix t assume t:"t\{x<..a" "tr x by (meson greaterThanLessThan_iff less_trans)+ have False when "f t = 0" using noroot t that by auto moreover have False when "f r=0" using noroot t that by auto moreover have False when "f r>0 \ f t<0 \ f r <0 \ f t >0" proof - have False when "{t..r} \ s \{}" proof - obtain t' where t':"t'\{t..r}" "t'\s" using \{t..r} \ s \ {}\ by blast then have "x t'r < b\ atLeastAtMost_iff greaterThanLessThan_iff less_le_trans not_le t) then have "t'\rs" unfolding rs_def using t \t'\s\ by auto then have "t'\Min rs" using \finite rs\ by auto moreover have "Min rs>r" using \r>x\ \t'\rs\ \finite rs\ unfolding r_def by (auto simp add:rs_def ) ultimately show False using t'(1) by auto qed moreover have "{t..r} \ {a<..x>t. x < r \ f x = 0" apply (intro IVT_strict) using that t assms(2) by auto then obtain t' where "t{x<..r}" unfolding rs_def using t by auto then show False using noroot \f t'=0\ by auto qed ultimately show "sgn (f t) = sgn (f r)" by (metis le_less not_less sgn_if) qed then show "f sgnx_able (at_right x)" by auto show "sgnx f (at_right x)\0" using noroot[of r,simplified] \(f has_sgnx sgn (f r)) (at_right x)\ by (simp add: has_sgnx_imp_sgnx sgn_if) qed lemma sgnx_able_poly[simp]: "(poly p) sgnx_able (at_right a)" "(poly p) sgnx_able (at_left a)" "(poly p) sgnx_able at_top" "(poly p) sgnx_able at_bot" proof - show "(poly p) sgnx_able at_top" using has_sgnx_def poly_sgn_eventually_at_top sgnx_able_def by blast show "(poly p) sgnx_able at_bot" using has_sgnx_def poly_sgn_eventually_at_bot sgnx_able_def by blast show "(poly p) sgnx_able (at_right a)" proof (cases "p=0") case True then show ?thesis unfolding sgnx_able_def has_sgnx_def eventually_at_right using linordered_field_no_ub by force next case False obtain ub where "ub>a" and ub:"\z. az\ub\poly p z\0" using next_non_root_interval[OF False] by auto have "\z. az\ub\sgn(poly p z) = sgn (poly p ub)" proof (rule ccontr) assume "\ (\z. a < z \ z \ ub \ sgn (poly p z) = sgn (poly p ub))" then obtain z where "aub" "sgn(poly p z) \ sgn (poly p ub)" by auto moreover then have "poly p z\0" "poly p ub\0" "z\ub" using ub \ub>a\ by blast+ ultimately have "(poly p z>0 \ poly p ub<0) \ (poly p z<0 \ poly p ub>0)" by (metis linorder_neqE_linordered_idom sgn_neg sgn_pos) then have "\x>z. x < ub \ poly p x = 0" using poly_IVT_neg[of z ub p] poly_IVT_pos[of z ub p] \z\ub\ \z\ub\ by argo then show False using ub \a < z\ by auto qed then show ?thesis unfolding sgnx_able_def has_sgnx_def eventually_at_right apply (rule_tac exI[where x="sgn(poly p ub)"]) apply (rule_tac exI[where x="ub"]) using less_eq_real_def \ub>a\ by blast qed show "(poly p) sgnx_able (at_left a)" proof (cases "p=0") case True then show ?thesis unfolding sgnx_able_def has_sgnx_def eventually_at_right using linordered_field_no_ub by force next case False obtain lb where "lbz. lb\z\zpoly p z\0" using last_non_root_interval[OF False] by auto have "\z. lb\z\zsgn(poly p z) = sgn (poly p lb)" proof (rule ccontr) assume "\ (\z. lb\z\z sgn (poly p z) = sgn (poly p lb))" then obtain z where "lb\z" "z sgn (poly p lb)" by auto moreover then have "poly p z\0" "poly p lb\0" "z\lb" using ub \lb by blast+ ultimately have "(poly p z>0 \ poly p lb<0) \ (poly p z<0 \ poly p lb>0)" by (metis linorder_neqE_linordered_idom sgn_neg sgn_pos) then have "\x>lb. x < z \ poly p x = 0" using poly_IVT_neg[of lb z p] poly_IVT_pos[of lb z p] \lb\z\ \z\lb\ by argo then show False using ub \z < a\ by auto qed then show ?thesis unfolding sgnx_able_def has_sgnx_def eventually_at_left apply (rule_tac exI[where x="sgn(poly p lb)"]) apply (rule_tac exI[where x="lb"]) using less_eq_real_def \lb by blast qed qed lemma has_sgnx_identity[intro,sgnx_intros]: shows "x\0 \((\x. x) has_sgnx 1) (at_right x)" "x\0 \ ((\x. x) has_sgnx -1) (at_left x)" proof - show "x\0 \ ((\x. x) has_sgnx 1) (at_right x)" unfolding has_sgnx_def eventually_at_right apply (intro exI[where x="x+1"]) by auto show "x\0 \ ((\x. x) has_sgnx -1) (at_left x)" unfolding has_sgnx_def eventually_at_left apply (intro exI[where x="x-1"]) by auto qed lemma has_sgnx_divide[sgnx_intros]: assumes "(f has_sgnx c1) F" "(g has_sgnx c2) F" shows "((\x. f x / g x) has_sgnx c1 / c2) F" proof - have "\\<^sub>F x in F. sgn (f x) = c1 \ sgn (g x) = c2" using assms unfolding has_sgnx_def by (intro eventually_conj,auto) then have "\\<^sub>F x in F. sgn (f x / g x) = c1 / c2" apply (elim eventually_mono) by (simp add: sgn_mult sgn_divide) then show "((\x. f x / g x) has_sgnx c1 / c2) F" unfolding has_sgnx_def by auto qed lemma sgnx_able_divide[sgnx_intros]: assumes "f sgnx_able F" "g sgnx_able F" shows "(\x. f x / g x) sgnx_able F" using has_sgnx_divide by (meson assms(1) assms(2) sgnx_able_def) lemma sgnx_divide: assumes "F\bot" "f sgnx_able F" "g sgnx_able F" shows "sgnx (\x. f x / g x) F =sgnx f F / sgnx g F" proof - obtain c1 c2 where c1:"(f has_sgnx c1) F" and c2:"(g has_sgnx c2) F" using assms unfolding sgnx_able_def by auto have "sgnx f F=c1" "sgnx g F=c2" using c1 c2 \F\bot\ by auto moreover have "((\x. f x / g x) has_sgnx c1 / c2) F" using has_sgnx_divide[OF c1 c2] . ultimately show ?thesis using assms(1) has_sgnx_imp_sgnx by blast qed lemma has_sgnx_times[sgnx_intros]: assumes "(f has_sgnx c1) F" "(g has_sgnx c2) F" shows "((\x. f x* g x) has_sgnx c1 * c2) F" proof - have "\\<^sub>F x in F. sgn (f x) = c1 \ sgn (g x) = c2" using assms unfolding has_sgnx_def by (intro eventually_conj,auto) then have "\\<^sub>F x in F. sgn (f x * g x) = c1 * c2" apply (elim eventually_mono) by (simp add: sgn_mult) then show "((\x. f x* g x) has_sgnx c1 * c2) F" unfolding has_sgnx_def by auto qed lemma sgnx_able_times[sgnx_intros]: assumes "f sgnx_able F" "g sgnx_able F" shows "(\x. f x * g x) sgnx_able F" using has_sgnx_times by (meson assms(1) assms(2) sgnx_able_def) lemma sgnx_times: assumes "F\bot" "f sgnx_able F" "g sgnx_able F" shows "sgnx (\x. f x * g x) F =sgnx f F * sgnx g F" proof - obtain c1 c2 where c1:"(f has_sgnx c1) F" and c2:"(g has_sgnx c2) F" using assms unfolding sgnx_able_def by auto have "sgnx f F=c1" "sgnx g F=c2" using c1 c2 \F\bot\ by auto moreover have "((\x. f x* g x) has_sgnx c1 * c2) F" using has_sgnx_times[OF c1 c2] . ultimately show ?thesis using assms(1) has_sgnx_imp_sgnx by blast qed lemma tendsto_nonzero_has_sgnx: assumes "(f \ c) F" "c\0" shows "(f has_sgnx sgn c) F" proof (cases rule:linorder_cases[of c 0]) case less then have "\\<^sub>F x in F. f x<0" using order_topology_class.order_tendstoD[OF assms(1),of 0] by auto then show ?thesis unfolding has_sgnx_def apply (elim eventually_mono) using less by auto next case equal then show ?thesis using \c\0\ by auto next case greater then have "\\<^sub>F x in F. f x>0" using order_topology_class.order_tendstoD[OF assms(1),of 0] by auto then show ?thesis unfolding has_sgnx_def apply (elim eventually_mono) using greater by auto qed lemma tendsto_nonzero_sgnx: assumes "(f \ c) F" "F\bot" "c\0" shows "sgnx f F = sgn c" using tendsto_nonzero_has_sgnx by (simp add: assms has_sgnx_imp_sgnx) lemma filterlim_divide_at_bot_at_top_iff: assumes "(f \ c) F" "c\0" shows "(LIM x F. f x / g x :> at_bot) \ (g \ 0) F \ ((\x. g x) has_sgnx - sgn c) F" "(LIM x F. f x / g x :> at_top) \ (g \ 0) F \ ((\x. g x) has_sgnx sgn c) F" proof - show "(LIM x F. f x / g x :> at_bot) \ ((g \ 0) F ) \ ((\x. g x) has_sgnx - sgn c) F" proof assume asm:"LIM x F. f x / g x :> at_bot" then have "filterlim g (at 0) F" using filterlim_at_infinity_divide_iff[OF assms(1,2),of g] at_bot_le_at_infinity filterlim_mono by blast then have "(g \ 0) F" using filterlim_at by blast moreover have "(g has_sgnx - sgn c) F" proof - have "((\x. sgn c * inverse (f x)) \ sgn c * inverse c) F" using assms(1,2) by (auto intro:tendsto_intros) then have "LIM x F. sgn c * inverse (f x) * (f x / g x) :> at_bot" apply (elim filterlim_tendsto_pos_mult_at_bot[OF _ _ asm]) using \c\0\ sgn_real_def by auto then have "LIM x F. sgn c / g x :> at_bot" apply (elim filterlim_mono_eventually) using eventually_times_inverse_1[OF assms] by (auto elim:eventually_mono) then have "\\<^sub>F x in F. sgn c / g x < 0" using filterlim_at_bot_dense[of "\x. sgn c/g x" F] by auto then show ?thesis unfolding has_sgnx_def apply (elim eventually_mono) by (metis add.inverse_inverse divide_less_0_iff sgn_neg sgn_pos sgn_sgn) qed ultimately show "(g \ 0) F \ (g has_sgnx - sgn c) F" by auto next assume "(g \ 0) F \ (g has_sgnx - sgn c) F" then have asm:"(g \ 0) F" "(g has_sgnx - sgn c) F" by auto have "LIM x F. inverse (g x * sgn c) :> at_bot" proof (rule filterlim_inverse_at_bot) show "((\x. g x * sgn c) \ 0) F" apply (rule tendsto_mult_left_zero) using asm(1) by blast next show "\\<^sub>F x in F. g x * sgn c < 0" using asm(2) unfolding has_sgnx_def apply (elim eventually_mono) by (metis add.inverse_inverse assms(2) linorder_neqE_linordered_idom mult_less_0_iff neg_0_less_iff_less sgn_greater sgn_zero_iff) qed moreover have "((\x. f x * sgn c) \ c * sgn c) F" using \(f \ c) F\ \c\0\ apply (intro tendsto_intros) by (auto simp add:sgn_zero_iff) moreover have "c * sgn c >0" using \c\0\ by (simp add: sgn_real_def) ultimately have "LIM x F. (f x * sgn c) *inverse (g x * sgn c) :> at_bot" using filterlim_tendsto_pos_mult_at_bot by blast then show "LIM x F. f x / g x :> at_bot" using \c\0\ by (auto simp add:field_simps sgn_zero_iff) qed show "(LIM x F. f x / g x :> at_top) \ ((g \ 0) F ) \ ((\x. g x) has_sgnx sgn c) F" proof assume asm:"LIM x F. f x / g x :> at_top" then have "filterlim g (at 0) F" using filterlim_at_infinity_divide_iff[OF assms(1,2),of g] at_top_le_at_infinity filterlim_mono by blast then have "(g \ 0) F" using filterlim_at by blast moreover have "(g has_sgnx sgn c) F" proof - have "((\x. sgn c * inverse (f x)) \ sgn c * inverse c) F" using assms(1,2) by (auto intro:tendsto_intros) then have "LIM x F. sgn c * inverse (f x) * (f x / g x) :> at_top" apply (elim filterlim_tendsto_pos_mult_at_top[OF _ _ asm]) using \c\0\ sgn_real_def by auto then have "LIM x F. sgn c / g x :> at_top" apply (elim filterlim_mono_eventually) using eventually_times_inverse_1[OF assms] by (auto elim:eventually_mono) then have "\\<^sub>F x in F. sgn c / g x > 0" using filterlim_at_top_dense[of "\x. sgn c/g x" F] by auto then show ?thesis unfolding has_sgnx_def apply (elim eventually_mono) by (metis sgn_greater sgn_less sgn_neg sgn_pos zero_less_divide_iff) qed ultimately show "(g \ 0) F \ (g has_sgnx sgn c) F" by auto next assume "(g \ 0) F \ (g has_sgnx sgn c) F" then have asm:"(g \ 0) F" "(g has_sgnx sgn c) F" by auto have "LIM x F. inverse (g x * sgn c) :> at_top" proof (rule filterlim_inverse_at_top) show "((\x. g x * sgn c) \ 0) F" apply (rule tendsto_mult_left_zero) using asm(1) by blast next show "\\<^sub>F x in F. g x * sgn c > 0" using asm(2) unfolding has_sgnx_def apply (elim eventually_mono) by (metis assms(2) sgn_1_neg sgn_greater sgn_if zero_less_mult_iff) qed moreover have "((\x. f x * sgn c) \ c * sgn c) F" using \(f \ c) F\ \c\0\ apply (intro tendsto_intros) by (auto simp add:sgn_zero_iff) moreover have "c * sgn c >0" using \c\0\ by (simp add: sgn_real_def) ultimately have "LIM x F. (f x * sgn c) *inverse (g x * sgn c) :> at_top" using filterlim_tendsto_pos_mult_at_top by blast then show "LIM x F. f x / g x :> at_top" using \c\0\ by (auto simp add:field_simps sgn_zero_iff) qed qed lemma poly_sgnx_left_right: fixes c a::real and p::"real poly" assumes "p\0" shows "sgnx (poly p) (at_left a) = (if even (order a p) then sgnx (poly p) (at_right a) else -sgnx (poly p) (at_right a))" using assms proof (induction "degree p" arbitrary: p rule: less_induct) case less have ?case when "poly p a\0" proof - have "sgnx (poly p) (at_left a) = sgn (poly p a)" by (simp add: has_sgnx_imp_sgnx tendsto_nonzero_has_sgnx that) moreover have "sgnx (poly p) (at_right a) = sgn (poly p a)" by (simp add: has_sgnx_imp_sgnx tendsto_nonzero_has_sgnx that) moreover have "order a p = 0" using that by (simp add: order_0I) ultimately show ?thesis by auto qed moreover have ?case when "poly p a=0" proof - obtain q where pq:"p= [:-a,1:] * q" using \poly p a=0\ by (meson dvdE poly_eq_0_iff_dvd) then have "q\0" using \p\0\ by auto then have "degree q < degree p" unfolding pq by (subst degree_mult_eq,auto) have "sgnx (poly p) (at_left a) = - sgnx (poly q) (at_left a)" proof - have "sgnx (\x. poly p x) (at_left a) = sgnx (poly q) (at_left a) * sgnx (poly [:-a,1:]) (at_left a)" unfolding pq apply (subst poly_mult) apply (subst sgnx_times) by auto moreover have "sgnx (\x. poly [:-a,1:] x) (at_left a) = -1" apply (intro has_sgnx_imp_sgnx) unfolding has_sgnx_def eventually_at_left by (auto simp add: linordered_field_no_lb) ultimately show ?thesis by auto qed moreover have "sgnx (poly p) (at_right a) = sgnx (poly q) (at_right a)" proof - have "sgnx (\x. poly p x) (at_right a) = sgnx (poly q) (at_right a) * sgnx (poly [:-a,1:]) (at_right a)" unfolding pq apply (subst poly_mult) apply (subst sgnx_times) by auto moreover have "sgnx (\x. poly [:-a,1:] x) (at_right a) = 1" apply (intro has_sgnx_imp_sgnx) unfolding has_sgnx_def eventually_at_right by (auto simp add: linordered_field_no_ub) ultimately show ?thesis by auto qed moreover have "even (order a p) \ odd (order a q)" unfolding pq apply (subst order_mult[OF \p \ 0\[unfolded pq]]) using \q\0\ by (auto simp add:order_power_n_n[of _ 1, simplified]) moreover note less.hyps[OF \degree q < degree p\ \q\0\] ultimately show ?thesis by auto qed ultimately show ?case by blast qed lemma poly_has_sgnx_left_right: fixes c a::real and p::"real poly" assumes "p\0" shows "(poly p has_sgnx c) (at_left a) \ (if even (order a p) then (poly p has_sgnx c) (at_right a) else (poly p has_sgnx -c) (at_right a))" using poly_sgnx_left_right by (metis (no_types, hide_lams) add.inverse_inverse assms has_sgnx_unique sgnx_able_poly sgnx_able_sgnx trivial_limit_at_left_real trivial_limit_at_right_real) lemma sign_r_pos_sgnx_iff: "sign_r_pos p a \ sgnx (poly p) (at_right a) > 0" proof assume asm:"0 < sgnx (poly p) (at_right a)" obtain c where c_def:"(poly p has_sgnx c) (at_right a)" using sgnx_able_poly(1) sgnx_able_sgnx by blast then have "c>0" using asm using has_sgnx_imp_sgnx trivial_limit_at_right_real by blast then show "sign_r_pos p a" using c_def unfolding sign_r_pos_def has_sgnx_def apply (elim eventually_mono) by force next assume asm:"sign_r_pos p a" define c where "c = sgnx (poly p) (at_right a)" then have "(poly p has_sgnx c) (at_right a)" by (simp add: sgnx_able_sgnx) then have "(\\<^sub>F x in (at_right a). poly p x>0 \ sgn (poly p x) = c)" using asm unfolding has_sgnx_def sign_r_pos_def by (simp add:eventually_conj_iff) then have "\\<^sub>F x in (at_right a). c > 0" apply (elim eventually_mono) by fastforce then show "c>0" by auto qed lemma sgnx_values: assumes "f sgnx_able F" "F \ bot" shows "sgnx f F = -1 \ sgnx f F = 0 \ sgnx f F = 1" proof - obtain c where c_def:"(f has_sgnx c) F" using assms(1) unfolding sgnx_able_def by auto then obtain x where "sgn(f x) = c" unfolding has_sgnx_def using assms(2) eventually_happens by blast then have "c=-1 \ c=0 \ c=1" using sgn_if by metis moreover have "sgnx f F = c" using c_def by (simp add: assms(2) has_sgnx_imp_sgnx) ultimately show ?thesis by auto qed lemma has_sgnx_poly_at_top: "(poly p has_sgnx sgn_pos_inf p) at_top" using has_sgnx_def poly_sgn_eventually_at_top by blast lemma has_sgnx_poly_at_bot: "(poly p has_sgnx sgn_neg_inf p) at_bot" using has_sgnx_def poly_sgn_eventually_at_bot by blast lemma sgnx_poly_at_top: "sgnx (poly p) at_top = sgn_pos_inf p" by (simp add: has_sgnx_def has_sgnx_imp_sgnx poly_sgn_eventually_at_top) lemma sgnx_poly_at_bot: "sgnx (poly p) at_bot = sgn_neg_inf p" by (simp add: has_sgnx_def has_sgnx_imp_sgnx poly_sgn_eventually_at_bot) lemma poly_has_sgnx_values: assumes "p\0" shows "(poly p has_sgnx 1) (at_left a) \ (poly p has_sgnx - 1) (at_left a)" "(poly p has_sgnx 1) (at_right a) \ (poly p has_sgnx - 1) (at_right a)" "(poly p has_sgnx 1) at_top \ (poly p has_sgnx - 1) at_top" "(poly p has_sgnx 1) at_bot \ (poly p has_sgnx - 1) at_bot" proof - have "sgn_pos_inf p = 1 \ sgn_pos_inf p = -1" unfolding sgn_pos_inf_def by (simp add: assms sgn_if) then show "(poly p has_sgnx 1) at_top \ (poly p has_sgnx - 1) at_top" using has_sgnx_poly_at_top by metis next have "sgn_neg_inf p = 1 \ sgn_neg_inf p = -1" unfolding sgn_neg_inf_def by (simp add: assms sgn_if) then show "(poly p has_sgnx 1) at_bot \ (poly p has_sgnx - 1) at_bot" using has_sgnx_poly_at_bot by metis next obtain c where c_def:"(poly p has_sgnx c) (at_left a)" using sgnx_able_poly(2) sgnx_able_sgnx by blast then have "sgnx (poly p) (at_left a) = c" using assms by auto then have "c=-1 \ c=0 \ c=1" using sgnx_values sgnx_able_poly(2) trivial_limit_at_left_real by blast moreover have False when "c=0" proof - have "(poly p has_sgnx 0) (at_left a)" using c_def that by auto then obtain lb where "lby. (lb y < a) \ poly p y = 0" unfolding has_sgnx_def eventually_at_left sgn_if by (metis one_neq_zero zero_neq_neg_one) then have "{lb<.. proots p" unfolding proots_within_def by auto then have "infinite (proots p)" apply (elim infinite_super) using \lb by auto moreover have "finite (proots p)" using finite_proots[OF \p\0\] by auto ultimately show False by auto qed ultimately have "c=-1 \ c=1" by auto then show "(poly p has_sgnx 1) (at_left a) \ (poly p has_sgnx - 1) (at_left a)" using c_def by auto next obtain c where c_def:"(poly p has_sgnx c) (at_right a)" using sgnx_able_poly(1) sgnx_able_sgnx by blast then have "sgnx (poly p) (at_right a) = c" using assms by auto then have "c=-1 \ c=0 \ c=1" using sgnx_values sgnx_able_poly(1) trivial_limit_at_right_real by blast moreover have False when "c=0" proof - have "(poly p has_sgnx 0) (at_right a)" using c_def that by auto then obtain ub where "ub>a" "\y. (a y < ub) \ poly p y = 0" unfolding has_sgnx_def eventually_at_right sgn_if by (metis one_neq_zero zero_neq_neg_one) then have "{a<.. proots p" unfolding proots_within_def by auto then have "infinite (proots p)" apply (elim infinite_super) using \ub>a\ by auto moreover have "finite (proots p)" using finite_proots[OF \p\0\] by auto ultimately show False by auto qed ultimately have "c=-1 \ c=1" by auto then show "(poly p has_sgnx 1) (at_right a) \ (poly p has_sgnx - 1) (at_right a)" using c_def by auto qed lemma poly_sgnx_values: assumes "p\0" shows "sgnx (poly p) (at_left a) = 1 \ sgnx (poly p) (at_left a) = -1" "sgnx (poly p) (at_right a) = 1 \ sgnx (poly p) (at_right a) = -1" using poly_has_sgnx_values[OF \p\0\] has_sgnx_imp_sgnx trivial_limit_at_left_real trivial_limit_at_right_real by blast+ lemma has_sgnx_inverse: "(f has_sgnx c) F \ ((inverse o f) has_sgnx (inverse c)) F" unfolding has_sgnx_def comp_def apply (rule eventually_subst) apply (rule always_eventually) by (metis inverse_inverse_eq sgn_inverse) lemma has_sgnx_derivative_at_left: assumes g_deriv:"(g has_field_derivative c) (at x)" and "g x=0" and "c\0" shows "(g has_sgnx - sgn c) (at_left x)" proof - have "(g has_sgnx -1) (at_left x)" when "c>0" proof - obtain d1 where "d1>0" and d1_def:"\h>0. h < d1 \ g (x - h) < g x" using DERIV_pos_inc_left[OF g_deriv \c>0\] \g x=0\ by auto have "(g has_sgnx -1) (at_left x)" unfolding has_sgnx_def eventually_at_left apply (intro exI[where x="x-d1"]) using \d1>0\ d1_def by (metis (no_types, hide_lams) add.commute add_uminus_conv_diff assms(2) diff_add_cancel diff_strict_left_mono diff_zero minus_diff_eq sgn_neg) thus ?thesis by auto qed moreover have "(g has_sgnx 1) (at_left x)" when "c<0" proof - obtain d1 where "d1>0" and d1_def:"\h>0. h < d1 \ g (x - h) > g x" using DERIV_neg_dec_left[OF g_deriv \c<0\] \g x=0\ by auto have "(g has_sgnx 1) (at_left x)" unfolding has_sgnx_def eventually_at_left apply (intro exI[where x="x-d1"]) using \d1>0\ d1_def by (metis (no_types, hide_lams) add.commute add_uminus_conv_diff assms(2) diff_add_cancel diff_zero less_diff_eq minus_diff_eq sgn_pos) thus ?thesis using \c<0\ by auto qed ultimately show ?thesis using \c\0\ using sgn_real_def by auto qed lemma has_sgnx_derivative_at_right: assumes g_deriv:"(g has_field_derivative c) (at x)" and "g x=0" and "c\0" shows "(g has_sgnx sgn c) (at_right x)" proof - have "(g has_sgnx 1) (at_right x)" when "c>0" proof - obtain d2 where "d2>0" and d2_def:"\h>0. h < d2 \ g x < g (x + h)" using DERIV_pos_inc_right[OF g_deriv \c>0\] \g x=0\ by auto have "(g has_sgnx 1) (at_right x)" unfolding has_sgnx_def eventually_at_right apply (intro exI[where x="x+d2"]) using \d2>0\ d2_def by (metis add.commute assms(2) diff_add_cancel diff_less_eq less_add_same_cancel1 sgn_pos) thus ?thesis using \c>0\ by auto qed moreover have "(g has_sgnx -1) (at_right x)" when "c<0" proof - obtain d2 where "d2>0" and d2_def:"\h>0. h < d2 \ g x > g (x + h)" using DERIV_neg_dec_right[OF g_deriv \c<0\] \g x=0\ by auto have "(g has_sgnx -1) (at_right x)" unfolding has_sgnx_def eventually_at_right apply (intro exI[where x="x+d2"]) using \d2>0\ d2_def by (metis (no_types, hide_lams) add.commute add.right_inverse add_uminus_conv_diff assms(2) diff_add_cancel diff_less_eq sgn_neg) thus ?thesis using \c<0\ by auto qed ultimately show ?thesis using \c\0\ using sgn_real_def by auto qed lemma has_sgnx_split: "(f has_sgnx c) (at x) \ (f has_sgnx c) (at_left x) \ (f has_sgnx c) (at_right x)" unfolding has_sgnx_def using eventually_at_split by auto lemma sgnx_at_top_IVT: assumes "sgnx (poly p) (at_right a) \ sgnx (poly p) at_top" shows "\x>a. poly p x=0" proof (cases "p=0") case True then show ?thesis using gt_ex[of a] by simp next case False from poly_has_sgnx_values[OF this] have "(poly p has_sgnx 1) (at_right a) \ (poly p has_sgnx - 1) (at_right a)" "(poly p has_sgnx 1) at_top \ (poly p has_sgnx - 1) at_top" by auto moreover have ?thesis when has_r:"(poly p has_sgnx 1) (at_right a)" and has_top:"(poly p has_sgnx -1) at_top" proof - obtain b where "b>a" "poly p b>0" proof - obtain a' where "a'>a" and a'_def:"\y>a. y < a' \ sgn (poly p y) = 1" using has_r[unfolded has_sgnx_def eventually_at_right] by auto define b where "b=(a+a')/2" have "aa'>a\ by auto moreover have "poly p b>0" using a'_def[rule_format,OF \b>a\ \b] unfolding sgn_if by argo ultimately show ?thesis using that by auto qed moreover obtain c where "c>b" "poly p c<0" proof - obtain b' where b'_def:"\n\b'. sgn (poly p n) = - 1" using has_top[unfolded has_sgnx_def eventually_at_top_linorder] by auto define c where "c=1+max b b'" have "c>b" "c\b'" unfolding c_def using \b>a\ by auto moreover have "poly p c<0" using b'_def[rule_format,OF \b'\c\] unfolding sgn_if by argo ultimately show ?thesis using that by auto qed ultimately show ?thesis using poly_IVT_neg[of b c p] not_less by fastforce qed moreover have ?thesis when has_r:"(poly p has_sgnx -1) (at_right a)" and has_top:"(poly p has_sgnx 1) at_top" proof - obtain b where "b>a" "poly p b<0" proof - obtain a' where "a'>a" and a'_def:"\y>a. y < a' \ sgn (poly p y) = -1" using has_r[unfolded has_sgnx_def eventually_at_right] by auto define b where "b=(a+a')/2" have "aa'>a\ by auto moreover have "poly p b<0" using a'_def[rule_format,OF \b>a\ \b] unfolding sgn_if by argo ultimately show ?thesis using that by auto qed moreover obtain c where "c>b" "poly p c>0" proof - obtain b' where b'_def:"\n\b'. sgn (poly p n) = 1" using has_top[unfolded has_sgnx_def eventually_at_top_linorder] by auto define c where "c=1+max b b'" have "c>b" "c\b'" unfolding c_def using \b>a\ by auto moreover have "poly p c>0" using b'_def[rule_format,OF \b'\c\] unfolding sgn_if by argo ultimately show ?thesis using that by auto qed ultimately show ?thesis using poly_IVT_pos[of b c p] not_less by fastforce qed moreover have ?thesis when "(poly p has_sgnx 1) (at_right a) \ (poly p has_sgnx 1) at_top \ (poly p has_sgnx - 1) (at_right a) \ (poly p has_sgnx -1) at_top" proof - have "sgnx (poly p) (at_right a) = sgnx (poly p) at_top" using that has_sgnx_imp_sgnx by auto then have False using assms by simp then show ?thesis by auto qed ultimately show ?thesis by blast qed lemma sgnx_at_left_at_right_IVT: assumes "sgnx (poly p) (at_right a) \ sgnx (poly p) (at_left b)" "ax. a x poly p x=0" proof (cases "p=0") case True then show ?thesis using \a by (auto intro:exI[where x="(a+b)/2"]) next case False from poly_has_sgnx_values[OF this] have "(poly p has_sgnx 1) (at_right a) \ (poly p has_sgnx - 1) (at_right a)" "(poly p has_sgnx 1) (at_left b) \ (poly p has_sgnx - 1) (at_left b)" by auto moreover have ?thesis when has_r:"(poly p has_sgnx 1) (at_right a)" and has_l:"(poly p has_sgnx -1) (at_left b)" proof - obtain c where "a0" proof - obtain a' where "a'>a" and a'_def:"\y>a. y < a' \ sgn (poly p y) = 1" using has_r[unfolded has_sgnx_def eventually_at_right] by auto define c where "c=(a+min a' b)/2" have "aa'>a\ \b>a\ by auto moreover have "poly p c>0" using a'_def[rule_format,OF \c>a\ \c] unfolding sgn_if by argo ultimately show ?thesis using that by auto qed moreover obtain d where "cy>b'. y < b \ sgn (poly p y) = - 1" using has_l[unfolded has_sgnx_def eventually_at_left] by auto define d where "d=(b+max b' c)/2" have "b'c" unfolding d_def using \b>b'\ \b>c\ by auto moreover have "poly p d<0" using b'_def[rule_format, OF \b' \d] unfolding sgn_if by argo ultimately show ?thesis using that by auto qed ultimately obtain x where "cc>a\ \d by (auto intro: exI[where x=x]) qed moreover have ?thesis when has_r:"(poly p has_sgnx -1) (at_right a)" and has_l:"(poly p has_sgnx 1) (at_left b)" proof - obtain c where "aa" and a'_def:"\y>a. y < a' \ sgn (poly p y) = -1" using has_r[unfolded has_sgnx_def eventually_at_right] by auto define c where "c=(a+min a' b)/2" have "aa'>a\ \b>a\ by auto moreover have "poly p c<0" using a'_def[rule_format,OF \c>a\ \c] unfolding sgn_if by argo ultimately show ?thesis using that by auto qed moreover obtain d where "c0" proof - obtain b' where "b'y>b'. y < b \ sgn (poly p y) = 1" using has_l[unfolded has_sgnx_def eventually_at_left] by auto define d where "d=(b+max b' c)/2" have "b'c" unfolding d_def using \b>b'\ \b>c\ by auto moreover have "poly p d>0" using b'_def[rule_format, OF \b' \d] unfolding sgn_if by argo ultimately show ?thesis using that by auto qed ultimately obtain x where "cc>a\ \d by (auto intro: exI[where x=x]) qed moreover have ?thesis when "(poly p has_sgnx 1) (at_right a) \ (poly p has_sgnx 1) (at_left b) \ (poly p has_sgnx - 1) (at_right a) \ (poly p has_sgnx -1) (at_left b)" proof - have "sgnx (poly p) (at_right a) = sgnx (poly p) (at_left b)" using that has_sgnx_imp_sgnx by auto then have False using assms by simp then show ?thesis by auto qed ultimately show ?thesis by blast qed lemma sgnx_at_bot_IVT: assumes "sgnx (poly p) (at_left a) \ sgnx (poly p) at_bot" shows "\x (poly p has_sgnx - 1) (at_left a)" "(poly p has_sgnx 1) at_bot \ (poly p has_sgnx - 1) at_bot" by auto moreover have ?thesis when has_l:"(poly p has_sgnx 1) (at_left a)" and has_bot:"(poly p has_sgnx -1) at_bot" proof - obtain b where "b0" proof - obtain a' where "a'y>a'. y < a \ sgn (poly p y) = 1" using has_l[unfolded has_sgnx_def eventually_at_left] by auto define b where "b=(a+a')/2" have "a>b" "b>a'" unfolding b_def using \a' by auto moreover have "poly p b>0" using a'_def[rule_format,OF \b>a'\ \b] unfolding sgn_if by argo ultimately show ?thesis using that by auto qed moreover obtain c where "cn\b'. sgn (poly p n) = - 1" using has_bot[unfolded has_sgnx_def eventually_at_bot_linorder] by auto define c where "c=min b b'- 1" have "cb'" unfolding c_def using \b by auto moreover have "poly p c<0" using b'_def[rule_format,OF \b'\c\] unfolding sgn_if by argo ultimately show ?thesis using that by auto qed ultimately show ?thesis using poly_IVT_pos[of c b p] using not_less by fastforce qed moreover have ?thesis when has_l:"(poly p has_sgnx -1) (at_left a)" and has_bot:"(poly p has_sgnx 1) at_bot" proof - obtain b where "by>a'. y < a \ sgn (poly p y) = -1" using has_l[unfolded has_sgnx_def eventually_at_left] by auto define b where "b=(a+a')/2" have "a>b" "b>a'" unfolding b_def using \a' by auto moreover have "poly p b<0" using a'_def[rule_format,OF \b>a'\ \b] unfolding sgn_if by argo ultimately show ?thesis using that by auto qed moreover obtain c where "c0" proof - obtain b' where b'_def:"\n\b'. sgn (poly p n) = 1" using has_bot[unfolded has_sgnx_def eventually_at_bot_linorder] by auto define c where "c=min b b'- 1" have "cb'" unfolding c_def using \b by auto moreover have "poly p c>0" using b'_def[rule_format,OF \b'\c\] unfolding sgn_if by argo ultimately show ?thesis using that by auto qed ultimately show ?thesis using poly_IVT_neg[of c b p] using not_less by fastforce qed moreover have ?thesis when "(poly p has_sgnx 1) (at_left a) \ (poly p has_sgnx 1) at_bot \ (poly p has_sgnx - 1) (at_left a) \ (poly p has_sgnx -1) at_bot" proof - have "sgnx (poly p) (at_left a) = sgnx (poly p) at_bot" using that has_sgnx_imp_sgnx by auto then have False using assms by simp then show ?thesis by auto qed ultimately show ?thesis by blast qed lemma sgnx_poly_nz: assumes "poly p x\0" shows "sgnx (poly p) (at_left x) = sgn (poly p x)" "sgnx (poly p) (at_right x) = sgn (poly p x)" proof - have "(poly p has_sgnx sgn(poly p x)) (at x)" apply (rule tendsto_nonzero_has_sgnx) using assms by auto then show "sgnx (poly p) (at_left x) = sgn (poly p x)" "sgnx (poly p) (at_right x) = sgn (poly p x)" unfolding has_sgnx_split by auto qed subsection \Finite predicate segments over an interval\ inductive finite_Psegments::"(real \ bool) \ real \ real \ bool" for P where emptyI: "a\b \ finite_Psegments P a b"| insertI_1: "\s\{a..P s;\t\{s<.. \ finite_Psegments P a b"| insertI_2: "\s\{a..P s;(\t\{s<..P t);finite_Psegments P a s\ \ finite_Psegments P a b" lemma finite_Psegments_pos_linear: assumes "finite_Psegments P (b*lb+c) (b*ub+c) " and "b>0" shows "finite_Psegments (P o (\t. b*t+c)) lb ub" proof - have [simp]:"b\0" using \b>0\ by auto show ?thesis proof (rule finite_Psegments.induct[OF assms(1), of "\lb' ub'. finite_Psegments (P o (\t. b*t+c)) ((lb'-c)/b) ((ub'-c)/b)",simplified]) (*really weird application of the induction rule, is there an alternative?*) fix lb ub f assume "(lb::real)\ub" then have "(lb - c) / b \ (ub - c) / b" using \b>0\ by (auto simp add:field_simps) then show "finite_Psegments (f \ (\t. b * t + c)) ((ub - c) / b) ((lb - c) / b)" by (rule finite_Psegments.emptyI) next fix s lb ub P assume asm: "lb \ s \ s < ub" "\t\{s<.. (\t. b * t + c)) ((lb - c) / b) ((s - c) / b)" "s = lb \ P s" show "finite_Psegments (P \ (\t. b * t + c)) ((lb - c) / b) ((ub - c) / b)" apply (rule finite_Psegments.insertI_1[of "(s-c)/b"]) using asm \b>0\ by (auto simp add:field_simps) next fix s lb ub P assume asm: "lb \ s \ s < ub" "\t\{s<.. P t" "finite_Psegments (P \ (\t. b * t + c)) ((lb - c) / b) ((s - c) / b)" "s=lb \ P s" show "finite_Psegments (P \ (\t. b * t + c)) ((lb - c) / b) ((ub - c) / b)" apply (rule finite_Psegments.insertI_2[of "(s-c)/b"]) using asm \b>0\ by (auto simp add:field_simps) qed qed lemma finite_Psegments_congE: assumes "finite_Psegments Q lb ub" "\t. \lb \ Q t \ P t " shows "finite_Psegments P lb ub" using assms proof (induct rule:finite_Psegments.induct) case (emptyI a b) then show ?case using finite_Psegments.emptyI by auto next case (insertI_1 s a b) show ?case proof (rule finite_Psegments.insertI_1[of s]) have "P s" when "s\a" proof - have "s\{a<..s \ {a.. that by auto then show ?thesis using insertI_1 by auto qed then show "s = a \ P s" by auto next show "s \ {a..t\{s<..a" proof - have "s\{a<..s \ {a.. that by auto then show ?thesis using insertI_2 by auto qed then show "s = a \ P s" by auto next show "s \ {a..t\{s<.. P t" "finite_Psegments P a s" using insertI_2 by auto qed qed lemma finite_Psegments_constI: assumes "\t. \a \ P t = c" shows "finite_Psegments P a b" proof - have "finite_Psegments (\_. c) a b" proof - have ?thesis when "a\b" using that finite_Psegments.emptyI by auto moreover have ?thesis when "ac" apply (rule finite_Psegments.insertI_2[of a]) using that by (auto intro: finite_Psegments.emptyI) ultimately show ?thesis by argo qed then show ?thesis apply (elim finite_Psegments_congE) using assms by auto qed context begin private lemma finite_Psegments_less_eq1: assumes "finite_Psegments P a c" "b\c" shows "finite_Psegments P a b" using assms proof (induct arbitrary: b rule:finite_Psegments.induct) case (emptyI a c) then show ?case using finite_Psegments.emptyI by auto next case (insertI_1 s a c) have ?case when "b\s" using insertI_1 that by auto moreover have ?case when "b>s" proof - have "s \ {a..s \ {a.. \b \ c\ by auto moreover have "\t\{s<..\t\{s<.. that \b \ c\ by auto ultimately show ?case using finite_Psegments.insertI_1[OF _ _ _ \finite_Psegments P a s\] \ s = a \ P s\ by auto qed ultimately show ?case by fastforce next case (insertI_2 s a c) have ?case when "b\s" using insertI_2 that by auto moreover have ?case when "b>s" proof - have "s \ {a..s \ {a.. \b \ c\ by auto moreover have "\t\{s<.. P t" using \\t\{s<.. P t\ that \b \ c\ by auto ultimately show ?case using finite_Psegments.insertI_2[OF _ _ _ \finite_Psegments P a s\] \ s = a \ P s\ by auto qed ultimately show ?case by fastforce qed private lemma finite_Psegments_less_eq2: assumes "finite_Psegments P a c" "a\b" shows "finite_Psegments P b c" using assms proof (induct arbitrary: rule:finite_Psegments.induct) case (emptyI a c) then show ?case using finite_Psegments.emptyI by auto next case (insertI_1 s a c) have ?case when "s\b" proof - have "\t\{b<..b" apply (rule finite_Psegments.insertI_1[where s=s]) using insertI_1 that by auto ultimately show ?case by linarith next case (insertI_2 s a c) have ?case when "s\b" proof - have "\t\{b<.. P t" using insertI_2 that by auto then show ?thesis by (metis finite_Psegments_constI greaterThanLessThan_iff) qed moreover have ?case when "s>b" apply (rule finite_Psegments.insertI_2[where s=s]) using insertI_2 that by auto ultimately show ?case by linarith qed lemma finite_Psegments_included: assumes "finite_Psegments P a d" "a\b" "c\d" shows "finite_Psegments P b c" using finite_Psegments_less_eq2 finite_Psegments_less_eq1 assms by blast end lemma finite_Psegments_combine: assumes "finite_Psegments P a b" "finite_Psegments P b c" "b\{a..c}" "closed ({x. P x} \ {a..c})" shows "finite_Psegments P a c" using assms(2,1,3,4) proof (induct rule:finite_Psegments.induct) case (emptyI b c) then show ?case using finite_Psegments_included by auto next case (insertI_1 s b c) have "P s" proof - have "s {s..(s+c)/2}" have "closed S" proof - have "closed ({a. P a} \ {a..c})" using insertI_1(8) . moreover have "S = ({a. P a} \ {a..c}) \ {s..(s+c)/2}" using insertI_1(1,7) unfolding S_def by (auto simp add:field_simps) ultimately show ?thesis using closed_Int[of "{a. P a} \ {a..c}" "{s..(s+c)/2}"] by blast qed moreover have "\y\S. dist y s < e" when "e>0" for e proof - define y where "y = min ((s+c)/2) (e/2+s)" have "y\S" proof - have "y\{s..(s+c)/2}" unfolding y_def using \e>0\ \s by (auto simp add:min_mult_distrib_left algebra_simps) moreover have "P y" apply (rule insertI_1(3)[rule_format]) unfolding y_def using \e>0\ \s by (auto simp add:algebra_simps min_mult_distrib_left min_less_iff_disj) ultimately show ?thesis unfolding S_def by auto qed moreover have "dist y s e>0\ \s by (auto simp add:algebra_simps min_mult_distrib_left min_less_iff_disj dist_real_def) ultimately show ?thesis by auto qed ultimately have "s\S" using closed_approachable by auto then show ?thesis unfolding S_def by auto qed show ?case proof (rule finite_Psegments.insertI_1[of s]) show " s \ {a.. P s" "\t\{s<..P s\ by auto next have "closed ({a. P a} \ {a..s})" using closed_Int[OF \closed ({a. P a} \ {a..c})\,of "{a..s}",simplified] apply (elim arg_elim[of closed]) using \s \ {b.. \b \ {a..c}\ by auto then show "finite_Psegments P a s" using insertI_1 by auto qed next case (insertI_2 s b c) have ?case when "P s" proof (rule finite_Psegments.insertI_2[of s]) show "s \ {a.. P s" "\t\{s<.. P t" using that insertI_2 by auto next have "closed ({a. P a} \ {a..s})" using closed_Int[OF \closed ({a. P a} \ {a..c})\,of "{a..s}",simplified] apply (elim arg_elim[of closed]) using \s \ {b.. \b \ {a..c}\ by auto then show "finite_Psegments P a s" using insertI_2 by auto qed moreover have ?case when "\ P s" "s=b" using \finite_Psegments P a b\ proof (cases rule:finite_Psegments.cases) case emptyI then show ?thesis using insertI_2 that by (metis antisym_conv atLeastAtMost_iff finite_Psegments.insertI_2) next case (insertI_1 s0) have "P s" proof - have "s0 {(s0+s)/2..s}" have "closed S" using closed_Int[OF \closed ({a. P a} \ {a..c})\,of "{(s0+s)/2..s}",simplified] apply (elim arg_elim[of closed]) unfolding S_def using \s0 \ {a.. \ s \ {b.. \b \ {a..c}\ by auto moreover have "\y\S. dist y s < e" when "e>0" for e proof - define y where "y = max ((s+s0)/2) (s-e/2)" have "y\S" proof - have "y\{(s0+s)/2..s}" unfolding y_def using \e>0\ \s0 by (auto simp add:field_simps min_mult_distrib_left) moreover have "P y" apply (rule insertI_1(3)[rule_format]) unfolding y_def using \e>0\ \s0 \s=b\ by (auto simp add:field_simps max_mult_distrib_left less_max_iff_disj) ultimately show ?thesis unfolding S_def by auto qed moreover have "dist y s e>0\ \s0 by (auto simp add:algebra_simps max_mult_distrib_left less_max_iff_disj dist_real_def max_add_distrib_right) ultimately show ?thesis by auto qed ultimately have "s\S" using closed_approachable by auto then show ?thesis unfolding S_def by auto qed then have False using \\ P s\ by auto then show ?thesis by simp next case (insertI_2 s0) have *: "\t\{s0<.. P t" using \ \t\{s<.. P t\ that \\t\{s0<.. P t\ by force show ?thesis apply (rule finite_Psegments.insertI_2[of s0]) subgoal using insertI_2.prems(2) local.insertI_2(1) by auto subgoal using \s0 = a \ P s0\ . subgoal using * . subgoal using \finite_Psegments P a s0\ . done qed moreover note \s = b \ P s\ ultimately show ?case by auto qed subsection \Finite segment intersection of a path with the imaginary axis\ definition finite_ReZ_segments::"(real \ complex) \ complex \ bool" where "finite_ReZ_segments g z = finite_Psegments (\t. Re (g t - z) = 0) 0 1" lemma finite_ReZ_segments_joinpaths: assumes g1:"finite_ReZ_segments g1 z" and g2: "finite_ReZ_segments g2 z" and "path g1" "path g2" "pathfinish g1=pathstart g2" shows "finite_ReZ_segments (g1+++g2) z" proof - define P where "P = (\t. (Re ((g1 +++ g2) t - z) = 0 \ 0 t<1) \ t=0 \ t=1)" have "finite_Psegments P 0 (1/2)" proof - have "finite_Psegments (\t. Re (g1 t - z) = 0) 0 1" using g1 unfolding finite_ReZ_segments_def . then have "finite_Psegments (\t. Re (g1 (2 * t) - z) = 0) 0 (1/2)" apply (drule_tac finite_Psegments_pos_linear[of _ 2 0 0 "1/2",simplified]) by (auto simp add:comp_def) then show ?thesis unfolding P_def joinpaths_def by (elim finite_Psegments_congE,auto) qed moreover have "finite_Psegments P (1/2) 1" proof - have "finite_Psegments (\t. Re (g2 t - z) = 0) 0 1" using g2 unfolding finite_ReZ_segments_def . then have "finite_Psegments (\t. Re (g2 (2 * t-1) - z) = 0) (1/2) 1" apply (drule_tac finite_Psegments_pos_linear[of _ 2 "1/2" "-1" 1,simplified]) by (auto simp add:comp_def) then show ?thesis unfolding P_def joinpaths_def apply (elim finite_Psegments_congE) by auto qed moreover have "closed {x. P x}" proof - define Q where "Q=(\t. Re ((g1 +++ g2) t - z) = 0)" have "continuous_on {0<..<1} (g1+++g2)" using path_join_imp[OF \path g1\ \path g2\ \pathfinish g1=pathstart g2\] unfolding path_def by (auto elim:continuous_on_subset) from continuous_on_Re[OF this] have "continuous_on {0<..<1} (\x. Re ((g1 +++ g2) x))" . from continuous_on_open_Collect_neq[OF this,of "\_. Re z",OF continuous_on_const,simplified] have "open {t. Re ((g1 +++ g2) t - z) \ 0 \ 0 t<1}" by (elim arg_elim[where f="open"],auto) from closed_Diff[of "{0::real..1}",OF _ this,simplified] show "closed {x. P x}" apply (elim arg_elim[where f=closed]) by (auto simp add:P_def) qed ultimately have "finite_Psegments P 0 1" using finite_Psegments_combine[of _ 0 "1/2" 1] by auto then show ?thesis unfolding finite_ReZ_segments_def P_def by (elim finite_Psegments_congE,auto) qed lemma finite_ReZ_segments_congE: assumes "finite_ReZ_segments p1 z1" "\t. \0 \ Re(p1 t- z1) = Re(p2 t - z2)" shows "finite_ReZ_segments p2 z2" using assms unfolding finite_ReZ_segments_def apply (elim finite_Psegments_congE) by auto lemma finite_ReZ_segments_constI: assumes "\t. 0t<1 \ g t = c" shows "finite_ReZ_segments g z" proof - have "finite_ReZ_segments (\_. c) z" unfolding finite_ReZ_segments_def by (rule finite_Psegments_constI,auto) then show ?thesis using assms by (elim finite_ReZ_segments_congE,auto) qed lemma finite_ReZ_segment_cases [consumes 1, case_names subEq subNEq,cases pred:finite_ReZ_segments]: assumes "finite_ReZ_segments g z" and subEq:"(\s. \s \ {0..<1};s=0\Re (g s) = Re z; \t\{s<..<1}. Re (g t) = Re z;finite_ReZ_segments (subpath 0 s g) z\ \ P)" and subNEq:"(\s. \s \ {0..<1};s=0\Re (g s) = Re z; \t\{s<..<1}. Re (g t) \ Re z;finite_ReZ_segments (subpath 0 s g) z\ \ P)" shows "P" using assms(1) unfolding finite_ReZ_segments_def proof (cases rule:finite_Psegments.cases) case emptyI then show ?thesis by auto next case (insertI_1 s) have "finite_ReZ_segments (subpath 0 s g) z" proof (cases "s=0") case True show ?thesis apply (rule finite_ReZ_segments_constI) using True unfolding subpath_def by auto next case False then have "s>0" using \s\{0..<1}\ by auto from finite_Psegments_pos_linear[OF _ this,of _ 0 0 1] insertI_1(4) show "finite_ReZ_segments (subpath 0 s g) z" unfolding finite_ReZ_segments_def comp_def subpath_def by auto qed then show ?thesis using subEq insertI_1 by force next case (insertI_2 s) have "finite_ReZ_segments (subpath 0 s g) z" proof (cases "s=0") case True show ?thesis apply (rule finite_ReZ_segments_constI) using True unfolding subpath_def by auto next case False then have "s>0" using \s\{0..<1}\ by auto from finite_Psegments_pos_linear[OF _ this,of _ 0 0 1] insertI_2(4) show "finite_ReZ_segments (subpath 0 s g) z" unfolding finite_ReZ_segments_def comp_def subpath_def by auto qed then show ?thesis using subNEq insertI_2 by force qed lemma finite_ReZ_segments_induct [case_names sub0 subEq subNEq, induct pred:finite_ReZ_segments]: assumes "finite_ReZ_segments g z" assumes sub0:"\g z. (P (subpath 0 0 g) z)" and subEq:"(\s g z. \s \ {0..<1};s=0\Re (g s) = Re z; \t\{s<..<1}. Re (g t) = Re z;finite_ReZ_segments (subpath 0 s g) z; P (subpath 0 s g) z\ \ P g z)" and subNEq:"(\s g z. \s \ {0..<1};s=0\Re (g s) = Re z; \t\{s<..<1}. Re (g t) \ Re z;finite_ReZ_segments (subpath 0 s g) z; P (subpath 0 s g) z\ \ P g z)" shows "P g z" proof - have "finite_Psegments (\t. Re (g t - z) = 0) 0 1" using assms(1) unfolding finite_ReZ_segments_def by auto then have "(0::real)\1 \ P (subpath 0 1 g) z" proof (induct rule: finite_Psegments.induct[of _ 0 1 "\a b. b\a \ P (subpath a b g) z"] ) case (emptyI a b) then show ?case using sub0[of "subpath a b g"] unfolding subpath_def by auto next case (insertI_1 s a b) have ?case when "a=b" using sub0[of "subpath a b g"] that unfolding subpath_def by auto moreover have ?case when "a\b" proof - have "b>a" using that \s \ {a.. by auto define s'::real where "s'=(s-a)/(b-a)" have "P (subpath a b g) z" proof (rule subEq[of s' "subpath a b g"]) show "\t\{s'<..<1}. Re (subpath a b g t) = Re z" proof fix t assume "t \ {s'<..<1}" then have "(b - a) * t + a\{s<..b>a\ \s \ {a.. apply (auto simp add:field_simps) by (sos "((((A<0 * (A<1 * A<2)) * R<1) + (((A<=1 * (A<0 * R<1)) * (R<1 * [1]^2)) + ((A<=0 * (A<0 * (A<1 * R<1))) * (R<1 * [1]^2)))))") then have "Re (g ((b - a) * t + a) - z) = 0" using insertI_1(3)[rule_format,of "(b - a) * t + a"] by auto then show "Re (subpath a b g t) = Re z" unfolding subpath_def by auto qed show "finite_ReZ_segments (subpath 0 s' (subpath a b g)) z" proof (cases "s=a") case True then show ?thesis unfolding s'_def subpath_def by (auto intro:finite_ReZ_segments_constI) next case False have "finite_Psegments (\t. Re (g t - z) = 0) a s" using insertI_1(4) unfolding finite_ReZ_segments_def by auto then have "finite_Psegments ((\t. Re (g t - z) = 0) \ (\t. (s - a) * t + a)) 0 1" apply (elim finite_Psegments_pos_linear[of _ "s-a" 0 a 1,simplified]) using False \s\{a.. by auto then show ?thesis using \b>a\ unfolding finite_ReZ_segments_def subpath_def s'_def comp_def by auto qed show "s' \ {0..<1}" using \b>a\ \s\{a.. unfolding s'_def by (auto simp add:field_simps) show "P (subpath 0 s' (subpath a b g)) z" proof - have "P (subpath a s g) z" using insertI_1(1,5) by auto then show ?thesis using \b>a\ unfolding s'_def subpath_def by simp qed show "s' = 0 \ Re (subpath a b g s') = Re z" proof - have ?thesis when "s=a" using that unfolding s'_def by auto moreover have ?thesis when "Re (g s - z) = 0" using that unfolding s'_def subpath_def by auto ultimately show ?thesis using \s = a \ Re (g s - z) = 0\ by auto qed qed then show ?thesis using \b>a\ by auto qed ultimately show ?case by auto next case (insertI_2 s a b) have ?case when "a=b" using sub0[of "subpath a b g"] that unfolding subpath_def by auto moreover have ?case when "a\b" proof - have "b>a" using that \s \ {a.. by auto define s'::real where "s'=(s-a)/(b-a)" have "P (subpath a b g) z" proof (rule subNEq[of s' "subpath a b g"]) show "\t\{s'<..<1}. Re (subpath a b g t) \ Re z" proof fix t assume "t \ {s'<..<1}" then have "(b - a) * t + a\{s<..b>a\ \s \ {a.. apply (auto simp add:field_simps) by (sos "((((A<0 * (A<1 * A<2)) * R<1) + (((A<=1 * (A<0 * R<1)) * (R<1 * [1]^2)) + ((A<=0 * (A<0 * (A<1 * R<1))) * (R<1 * [1]^2)))))") then have "Re (g ((b - a) * t + a) - z) \ 0" using insertI_2(3)[rule_format,of "(b - a) * t + a"] by auto then show "Re (subpath a b g t) \ Re z" unfolding subpath_def by auto qed show "finite_ReZ_segments (subpath 0 s' (subpath a b g)) z" proof (cases "s=a") case True then show ?thesis unfolding s'_def subpath_def by (auto intro:finite_ReZ_segments_constI) next case False have "finite_Psegments (\t. Re (g t - z) = 0) a s" using insertI_2(4) unfolding finite_ReZ_segments_def by auto then have "finite_Psegments ((\t. Re (g t - z) = 0) \ (\t. (s - a) * t + a)) 0 1" apply (elim finite_Psegments_pos_linear[of _ "s-a" 0 a 1,simplified]) using False \s\{a.. by auto then show ?thesis using \b>a\ unfolding finite_ReZ_segments_def subpath_def s'_def comp_def by auto qed show "s' \ {0..<1}" using \b>a\ \s\{a.. unfolding s'_def by (auto simp add:field_simps) show "P (subpath 0 s' (subpath a b g)) z" proof - have "P (subpath a s g) z" using insertI_2(1,5) by auto then show ?thesis using \b>a\ unfolding s'_def subpath_def by simp qed show "s' = 0 \ Re (subpath a b g s') = Re z" proof - have ?thesis when "s=a" using that unfolding s'_def by auto moreover have ?thesis when "Re (g s - z) = 0" using that unfolding s'_def subpath_def by auto ultimately show ?thesis using \s = a \ Re (g s - z) = 0\ by auto qed qed then show ?thesis using \b>a\ by auto qed ultimately show ?case by auto qed then show ?thesis by auto qed lemma finite_ReZ_segments_shiftpah: assumes "finite_ReZ_segments g z" "s\{0..1}" "path g" and loop:"pathfinish g = pathstart g" shows "finite_ReZ_segments (shiftpath s g) z" proof - have "finite_Psegments (\t. Re (shiftpath s g t - z) = 0) 0 (1-s)" proof - have "finite_Psegments (\t. Re (g t) = Re z) s 1" using assms finite_Psegments_included[of _ 0 1 s] unfolding finite_ReZ_segments_def by force then have "finite_Psegments (\t. Re (g (s + t) - z) = 0) 0 (1-s)" using finite_Psegments_pos_linear[of "\t. Re (g t - z) =0" 1 0 s "1-s",simplified] unfolding comp_def by (auto simp add:algebra_simps) then show ?thesis unfolding shiftpath_def apply (elim finite_Psegments_congE) using \s\{0..1}\ by auto qed moreover have "finite_Psegments (\t. Re (shiftpath s g t - z) = 0) (1-s) 1" proof - have "finite_Psegments (\t. Re (g t) = Re z) 0 s " using assms finite_Psegments_included unfolding finite_ReZ_segments_def by force then have "finite_Psegments (\t. Re (g (s + t - 1) - z) = 0) (1-s) 1" using finite_Psegments_pos_linear[of "\t. Re (g t - z) =0" 1 "1-s" "s-1" 1,simplified] unfolding comp_def by (auto simp add:algebra_simps) then show ?thesis unfolding shiftpath_def apply (elim finite_Psegments_congE) using \s\{0..1}\ by auto qed moreover have "1 - s \ {0..1}" using \s\{0..1}\ by auto moreover have "closed ({x. Re (shiftpath s g x - z) = 0} \ {0..1})" proof - let ?f = "\x. Re (shiftpath s g x - z)" have "continuous_on {0..1} ?f" using path_shiftpath[OF \path g\ loop \s\{0..1}\] unfolding path_def by (auto intro: continuous_intros) from continuous_closed_preimage_constant[OF this,of 0,simplified] show ?thesis apply (elim arg_elim[of closed]) by force qed ultimately show ?thesis unfolding finite_ReZ_segments_def by (rule finite_Psegments_combine[where b="1-s"]) qed lemma finite_imp_finite_ReZ_segments: assumes "finite {t. Re (g t - z) = 0 \ 0 \ t \ t\1}" shows "finite_ReZ_segments g z" proof - define P where "P = (\t. Re (g t - z) = 0)" define rs where "rs=(\b. {t. P t \ 0 < t \ t0" for b using that proof (induct "card (rs b)" arbitrary:b rule:nat_less_induct) case ind:1 have ?case when "rs b= {}" apply (rule finite_Psegments.intros(3)[of 0]) using that \0 < b\ unfolding rs_def by (auto intro:finite_Psegments.intros) moreover have ?case when "rs b\{}" proof - define lj where "lj = Max (rs b)" have "0finite (rs b)\ \rs b\{}\,folded lj_def] unfolding rs_def by auto show ?thesis proof (rule finite_Psegments.intros(3)[of lj]) show "lj \ {0.. P lj" using \0 \lj \P lj\ by auto show "\t\{lj<.. P t" proof (rule ccontr) assume " \ (\t\{lj<.. P t) " then obtain t where t:"P t" "lj < t" "t < b" by auto then have "t\rs b" unfolding rs_def using \lj>0\ by auto then have "t\lj" using Max_ge[OF \finite (rs b)\,of t] unfolding lj_def by auto then show False using \t>lj\ by auto qed show "finite_Psegments P 0 lj" proof (rule ind.hyps[rule_format,of "card (rs lj)" lj,simplified]) show "finite (rs lj)" using \finite (rs b)\ unfolding rs_def using \lj by (auto elim!:rev_finite_subset ) show "card (rs lj) < card (rs b)" apply (rule psubset_card_mono[OF \finite (rs b)\]) using Max_in \finite (rs lj)\ \lj < b\ lj_def rs_def that by fastforce show "0 < lj" using \0 . qed qed qed ultimately show ?case by auto qed moreover have "finite (rs 1)" using assms unfolding rs_def P_def by (auto elim:rev_finite_subset) ultimately have "finite_Psegments P 0 1" by auto then show ?thesis unfolding P_def finite_ReZ_segments_def . qed lemma finite_ReZ_segments_poly_linepath: shows "finite_ReZ_segments (poly p o linepath a b) z" proof - define P where "P=map_poly Re (pcompose (p-[:z:]) [:a,b-a:])" have *:"Re ((poly p \ linepath a b) t - z) = 0 \ poly P t=0" for t unfolding inner_complex_def P_def linepath_def comp_def apply (subst Re_poly_of_real[symmetric]) by (auto simp add: algebra_simps poly_pcompose scaleR_conv_of_real) have ?thesis when "P\0" proof - have "finite {t. poly P t=0}" using that poly_roots_finite by auto then have "finite {t. Re ((poly p \ linepath a b) t - z) = 0 \ 0 \ t \ t \ 1}" using * by auto then show ?thesis using finite_imp_finite_ReZ_segments[of "poly p o linepath a b" z] by auto qed moreover have ?thesis when "P=0" unfolding finite_ReZ_segments_def apply (rule finite_Psegments_constI[where c=True]) apply (subst *) using that by auto ultimately show ?thesis by auto qed lemma part_circlepath_half_finite_inter: assumes "st\tt" "r\0" "c\0" shows "finite {t. part_circlepath z0 r st tt t \ c = d \ 0\ t \ t\1}" (is "finite ?T") proof - let ?S = "{\. (z0+r*exp (\ * \ )) \ c = d \ \ \ closed_segment st tt}" define S where "S \ {\. (z0+r*exp (\ * \ )) \ c = d \ \ \ closed_segment st tt}" have "S = linepath st tt ` ?T" proof define g where "g\(\t. (t-st)/(tt -st))" have "0\g t" "g t\1" when "t \ closed_segment st tt " for t using that \st\tt\ closed_segment_eq_real_ivl unfolding g_def real_scaleR_def by (auto simp add:divide_simps) moreover have "linepath st tt (g t) =t" "g (linepath st tt t) = t" for t unfolding linepath_def g_def real_scaleR_def using \st\tt\ apply (simp_all add:divide_simps) by (auto simp add:algebra_simps ) ultimately have "x\linepath st tt ` ?T" when "x\S" for x using that unfolding S_def by (auto intro!:image_eqI[where x="g x"] simp add:part_circlepath_def) then show "S \ linepath st tt ` ?T" by auto next have "x\S" when "x\linepath st tt ` ?T" for x using that unfolding part_circlepath_def S_def by (auto simp add: linepath_in_path) then show "linepath st tt ` ?T \ S" by auto qed moreover have "finite S" proof - define a' b' c' where "a'=r * Re c" and "b' = r* Im c" and "c'=Im c * Im z0 + Re z0 * Re c - d" define f where "f \= a' * cos \ + b' * sin \ + c'" for \ have "(z0+r*exp (\ * \ )) \ c = d \ f \ = 0" for \ unfolding exp_Euler inner_complex_def f_def a'_def b'_def c'_def by (auto simp add:algebra_simps cos_of_real sin_of_real) then have *:"S = roots f \ closed_segment st tt" unfolding S_def roots_within_def by auto have "uniform_discrete S" proof - have "a' \ 0 \ b' \ 0 \ c' \ 0" using assms complex_eq_iff unfolding a'_def b'_def c'_def by auto then have "periodic_set (roots f) (4 * pi)" using periodic_set_sin_cos_linear[of a' b' c',folded f_def] by auto then have "uniform_discrete (roots f)" using periodic_imp_uniform_discrete by auto then show ?thesis unfolding * by auto qed moreover have "bounded S" unfolding * by (simp add: bounded_Int bounded_closed_segment) ultimately show ?thesis using uniform_discrete_finite_iff by auto qed moreover have "inj_on (linepath st tt) ?T" proof - have "inj (linepath st tt)" unfolding linepath_def using assms inj_segment by blast then show ?thesis by (auto elim:subset_inj_on) qed ultimately show ?thesis by (auto elim!: finite_imageD) qed lemma linepath_half_finite_inter: assumes "a \ c \ d \ b \ c \ d" shows "finite {t. linepath a b t \ c = d \ 0\ t \ t\1}" (is "finite ?S") proof (rule ccontr) assume asm:"infinite ?S" obtain t1 t2 where u1u2:"t1\t2" "t1\?S" "t2\?S" proof - obtain t1 where "t1\?S" using not_finite_existsD asm by blast moreover have "\u2. u2\?S-{t1}" using infinite_remove[OF asm,of t1] by (meson finite.emptyI rev_finite_subset subsetI) ultimately show ?thesis using that by auto qed have t1:"(1-t1)*(a \ c) + t1 * (b \ c) = d" using \t1\?S\ unfolding linepath_def by (simp add: inner_left_distrib) have t2:"(1-t2)*(a \ c) + t2 * (b \ c) = d" using \t2\?S\ unfolding linepath_def by (simp add: inner_left_distrib) have "a \ c = d" proof - have "t2*((1-t1)*(a \ c) + t1 * (b \ c)) = t2*d" using t1 by auto then have *:"(t2-t1*t2)*(a \ c) + t1*t2 * (b \ c) = t2*d" by (auto simp add:algebra_simps) have "t1*((1-t2)*(a \ c) + t2 * (b \ c)) = t1*d" using t2 by auto then have **:"(t1-t1*t2)*(a \ c) + t1*t2 * (b \ c) = t1*d" by (auto simp add:algebra_simps) have "(t2-t1)*(a \ c) = (t2-t1)*d" using arg_cong2[OF * **,of minus] by (auto simp add:algebra_simps) then show ?thesis using \t1\t2\ by auto qed moreover have "b \ c = d" proof - have "(1-t2)*((1-t1)*(a \ c) + t1 * (b \ c)) = (1-t2)*d" using t1 by auto then have *:"(1-t1)*(1-t2)*(a \ c) + (t1-t1*t2) * (b \ c) = (1-t2)*d" by (auto simp add:algebra_simps) have "(1-t1)*((1-t2)*(a \ c) + t2 * (b \ c)) = (1-t1)*d" using t2 by auto then have **:"(1-t1)*(1-t2)*(a \ c) + (t2-t1*t2) * (b \ c) = (1-t1)*d" by (auto simp add:algebra_simps) have "(t2-t1)*(b \ c) = (t2-t1)*d" using arg_cong2[OF ** *,of minus] by (auto simp add:algebra_simps) then show ?thesis using \t1\t2\ by auto qed ultimately show False using assms by auto qed lemma finite_half_joinpaths_inter: assumes "finite {t. l1 t \ c = d \ 0\ t \ t\1}" "finite {t. l2 t \ c = d \ 0\ t \ t\1}" shows "finite {t. (l1+++l2) t \ c = d \ 0\ t \ t\1}" proof - let ?l1s = "{t. l1 (2*t) \ c = d \ 0\ t \ t\1/2}" let ?l2s = "{t. l2 (2 * t - 1) \ c = d \ 1/2< t \ t\1}" let ?ls = "\l. {t. l t \ c = d \ 0\ t \ t\1}" have "{t. (l1+++l2) t \ c = d \ 0\ t \ t\1} = ?l1s \ ?l2s" unfolding joinpaths_def by auto moreover have "finite ?l1s" proof - have "?l1s = ((*) (1/2)) ` ?ls l1" by (auto intro:rev_image_eqI) thus ?thesis using assms by simp qed moreover have "finite ?l2s" proof - have "?l2s \ (\x. x/2 + 1/2) ` ?ls l2" by (auto intro:rev_image_eqI simp add:field_simps) thus ?thesis using assms by (auto elim:finite_subset) qed ultimately show ?thesis by simp qed lemma finite_ReZ_segments_linepath: "finite_ReZ_segments (linepath a b) z" proof - have ?thesis when "Re a\Re z \ Re b \Re z" proof - let ?S1="{t. Re (linepath a b t-z) = 0 \ 0 \ t \ t \ 1}" have "finite ?S1" using linepath_half_finite_inter[of a "Complex 1 0" "Re z" b] that using one_complex.code by auto from finite_imp_finite_ReZ_segments[OF this] show ?thesis . qed moreover have ?thesis when "Re a=Re z" "Re b=Re z" unfolding finite_ReZ_segments_def apply (rule finite_Psegments.intros(2)[of 0]) using that unfolding linepath_def by (auto simp add:algebra_simps intro:finite_Psegments.intros) ultimately show ?thesis by blast qed lemma finite_ReZ_segments_part_circlepath: "finite_ReZ_segments (part_circlepath z0 r st tt) z" proof - have ?thesis when "st \ tt" "r \ 0" proof - let ?S1="{t. Re (part_circlepath z0 r st tt t-z) = 0 \ 0 \ t \ t \ 1}" have "finite ?S1" using part_circlepath_half_finite_inter[of st tt r "Complex 1 0" z0 "Re z"] that one_complex.code by (auto simp add:inner_complex_def ) from finite_imp_finite_ReZ_segments[OF this] show ?thesis . qed moreover have ?thesis when "st =tt \ r=0" proof - define c where "c = z0 + r * exp (\ * tt)" have "part_circlepath z0 r st tt = (\t. c)" unfolding part_circlepath_def c_def using that linepath_refl by auto then show ?thesis using finite_ReZ_segments_linepath[of c c z] linepath_refl[of c] by auto qed ultimately show ?thesis by blast qed lemma finite_ReZ_segments_poly_of_real: shows "finite_ReZ_segments (poly p o of_real) z" using finite_ReZ_segments_poly_linepath[of p 0 1 z] unfolding linepath_def by (auto simp add:scaleR_conv_of_real) lemma finite_ReZ_segments_subpath: assumes "finite_ReZ_segments g z" "0\u" "u\v" "v\1" shows "finite_ReZ_segments (subpath u v g) z" proof (cases "u=v") case True then show ?thesis unfolding subpath_def by (auto intro:finite_ReZ_segments_constI) next case False then have "uu\v\ by auto define P where "P=(\t. Re (g t - z) = 0)" have "finite_ReZ_segments (subpath u v g) z = finite_Psegments (P o (\t. (v - u) * t + u)) 0 1" unfolding finite_ReZ_segments_def subpath_def P_def comp_def by auto also have "..." apply (rule finite_Psegments_pos_linear) using assms False unfolding finite_ReZ_segments_def by (fold P_def,auto elim:finite_Psegments_included) finally show ?thesis . qed subsection \jump and jumpF\ definition jump::"(real \ real) \ real \ int" where "jump f a = (if (LIM x (at_left a). f x :> at_bot) \ (LIM x (at_right a). f x :> at_top) then 1 else if (LIM x (at_left a). f x :> at_top) \ (LIM x (at_right a). f x :> at_bot) then -1 else 0)" definition jumpF::"(real \ real) \ real filter \ real" where "jumpF f F \ (if filterlim f at_top F then 1/2 else if filterlim f at_bot F then -1/2 else (0::real))" lemma jumpF_const[simp]: assumes "F\bot" shows "jumpF (\_. c) F = 0" proof - have False when "LIM x F. c :> at_bot" using filterlim_at_bot_nhds[OF that _ \F\bot\] by auto moreover have False when "LIM x F. c :> at_top" using filterlim_at_top_nhds[OF that _ \F\bot\] by auto ultimately show ?thesis unfolding jumpF_def by auto qed lemma jumpF_not_infinity: assumes "continuous F g" "F\bot" shows "jumpF g F = 0" proof - have "\ filterlim g at_infinity F" using not_tendsto_and_filterlim_at_infinity[OF \F \bot\ assms(1)[unfolded continuous_def]] by auto then have "\ filterlim g at_bot F" "\ filterlim g at_top F" using at_bot_le_at_infinity at_top_le_at_infinity filterlim_mono by blast+ then show ?thesis unfolding jumpF_def by auto qed lemma jumpF_linear_comp: assumes "c\0" shows "jumpF (f o (\x. c*x+b)) (at_left x) = (if c>0 then jumpF f (at_left (c*x+b)) else jumpF f (at_right (c*x+b)))" (is ?case1) "jumpF (f o (\x. c*x+b)) (at_right x) = (if c>0 then jumpF f (at_right (c*x+b)) else jumpF f (at_left (c*x+b)))" (is ?case2) proof - let ?g = "\x. c*x+b" have ?case1 ?case2 when "\ c>0" proof - have "c<0" using \c\0\ that by auto have "filtermap ?g (at_left x) = at_right (?g x)" "filtermap ?g (at_right x) = at_left (?g x)" using \c<0\ filtermap_linear_at_left[OF \c\0\, of b x] filtermap_linear_at_right[OF \c\0\, of b x] by auto then have "jumpF (f \ ?g) (at_left x) = jumpF f (at_right (?g x))" "jumpF (f \ ?g) (at_right x) = jumpF f (at_left (?g x))" unfolding jumpF_def filterlim_def comp_def by (auto simp add: filtermap_filtermap[of f ?g,symmetric]) then show ?case1 ?case2 using \c<0\ by auto qed moreover have ?case1 ?case2 when "c>0" proof - have "filtermap ?g (at_left x) = at_left (?g x)" "filtermap ?g (at_right x) = at_right (?g x)" using that filtermap_linear_at_left[OF \c\0\, of b x] filtermap_linear_at_right[OF \c\0\, of b x] by auto then have "jumpF (f \ ?g) (at_left x) = jumpF f (at_left (?g x))" "jumpF (f \ ?g) (at_right x) = jumpF f (at_right (?g x))" unfolding jumpF_def filterlim_def comp_def by (auto simp add: filtermap_filtermap[of f ?g,symmetric]) then show ?case1 ?case2 using that by auto qed ultimately show ?case1 ?case2 by auto qed lemma jump_const[simp]:"jump (\_. c) a = 0" proof - have False when "LIM x (at_left a). c :> at_bot" apply (rule not_tendsto_and_filterlim_at_infinity[of "at_left a" "\_. c" c]) apply auto using at_bot_le_at_infinity filterlim_mono that by blast moreover have False when "LIM x (at_left a). c :> at_top" apply (rule not_tendsto_and_filterlim_at_infinity[of "at_left a" "\_. c" c]) apply auto using at_top_le_at_infinity filterlim_mono that by blast ultimately show ?thesis unfolding jump_def by auto qed lemma jump_not_infinity: "isCont f a \ jump f a =0" by (meson at_bot_le_at_infinity at_top_le_at_infinity filterlim_at_split filterlim_def isCont_def jump_def not_tendsto_and_filterlim_at_infinity order_trans trivial_limit_at_left_real) lemma jump_jump_poly_aux: assumes "p\0" "coprime p q" shows "jump (\x. poly q x / poly p x) a = jump_poly q p a" proof (cases "q=0") case True then show ?thesis by auto next case False define f where "f \ (\x. poly q x / poly p x)" have ?thesis when "poly q a = 0" proof - have "poly p a\0" using coprime_poly_0[OF \coprime p q\] that by blast then have "isCont f a" unfolding f_def by simp then have "jump f a=0" using jump_not_infinity by auto moreover have "jump_poly q p a=0" using jump_poly_not_root[OF \poly p a\0\] by auto ultimately show ?thesis unfolding f_def by auto qed moreover have ?thesis when "poly q a\0" proof (cases "even(order a p)") case True define c where "c\sgn (poly q a)" note filterlim_divide_at_bot_at_top_iff [OF _ that,of "poly q" "at_left a" "poly p",folded f_def c_def,simplified] filterlim_divide_at_bot_at_top_iff [OF _ that,of "poly q" "at_right a" "poly p",folded f_def c_def,simplified] moreover have "(poly p has_sgnx - c) (at_left a) = (poly p has_sgnx - c) (at_right a)" "(poly p has_sgnx c) (at_left a) = (poly p has_sgnx c) (at_right a)" using poly_has_sgnx_left_right[OF \p\0\] True by auto moreover have "c\0" by (simp add: c_def sgn_if that) then have False when "(poly p has_sgnx - c) (at_right a)" "(poly p has_sgnx c) (at_right a)" using has_sgnx_unique[OF _ that] by auto ultimately have "jump f a = 0" unfolding jump_def by auto moreover have "jump_poly q p a = 0" unfolding jump_poly_def using True by (simp add: order_0I that) ultimately show ?thesis unfolding f_def by auto next case False define c where "c\sgn (poly q a)" have "(poly p \ 0) (at a)" using False by (metis even_zero order_0I poly_tendsto(1)) then have "(poly p \ 0) (at_left a)" and "(poly p \ 0) (at_right a)" by (auto simp add: filterlim_at_split) moreover note filterlim_divide_at_bot_at_top_iff [OF _ that,of "poly q" _ "poly p",folded f_def c_def] moreover have "(poly p has_sgnx c) (at_left a) = (poly p has_sgnx - c) (at_right a)" "(poly p has_sgnx - c) (at_left a) = (poly p has_sgnx c) (at_right a)" using poly_has_sgnx_left_right[OF \p\0\] False by auto ultimately have "jump f a = (if (poly p has_sgnx c) (at_right a) then 1 else if (poly p has_sgnx - c) (at_right a) then -1 else 0)" unfolding jump_def by auto also have "... = (if sign_r_pos (q * p) a then 1 else - 1)" proof - have "(poly p has_sgnx c) (at_right a) \ sign_r_pos (q * p) a " proof assume "(poly p has_sgnx c) (at_right a)" then have "sgnx (poly p) (at_right a) = c" by auto moreover have "sgnx (poly q) (at_right a) = c" unfolding c_def using that by (auto intro!: tendsto_nonzero_sgnx) ultimately have "sgnx (\x. poly (q*p) x) (at_right a) = c * c" by (simp add:sgnx_times) moreover have "c\0" by (simp add: c_def sgn_if that) ultimately have "sgnx (\x. poly (q*p) x) (at_right a) > 0" using not_real_square_gt_zero by fastforce then show "sign_r_pos (q * p) a" using sign_r_pos_sgnx_iff by blast next assume asm:"sign_r_pos (q * p) a" let ?c1 = "sgnx (poly p) (at_right a)" let ?c2 = "sgnx (poly q) (at_right a)" have "0 < sgnx (\x. poly (q * p) x) (at_right a)" using asm sign_r_pos_sgnx_iff by blast then have "?c2 * ?c1 >0" apply (subst (asm) poly_mult) apply (subst (asm) sgnx_times) by auto then have "?c2>0 \ ?c1>0 \ ?c2<0 \ ?c1<0" by (simp add: zero_less_mult_iff) then have "?c1=?c2" using sgnx_values[OF sgnx_able_poly(1), of a,simplified] by (metis add.inverse_neutral less_minus_iff less_not_sym) moreover have "sgnx (poly q) (at_right a) = c" unfolding c_def using that by (auto intro!: tendsto_nonzero_sgnx) ultimately have "?c1 = c" by auto then show "(poly p has_sgnx c) (at_right a)" using sgnx_able_poly(1) sgnx_able_sgnx by blast qed then show ?thesis unfolding jump_poly_def using poly_has_sgnx_values[OF \p\0\] by (metis add.inverse_inverse c_def sgn_if that) qed also have "... = jump_poly q p a" unfolding jump_poly_def using False order_root that by (simp add: order_root assms(1)) finally show ?thesis unfolding f_def by auto qed ultimately show ?thesis by auto qed lemma jump_jumpF: assumes cont:"isCont (inverse o f) a" and sgnxl:"(f has_sgnx l) (at_left a)" and sgnxr:"(f has_sgnx r) (at_right a)" and "l\0 " "r\0" shows "jump f a = jumpF f (at_right a) - jumpF f (at_left a)" proof - have ?thesis when "filterlim f at_bot (at_left a)" "filterlim f at_top (at_right a)" unfolding jump_def jumpF_def using that filterlim_at_top_at_bot[OF _ _ trivial_limit_at_left_real] by auto moreover have ?thesis when "filterlim f at_top (at_left a)" "filterlim f at_bot (at_right a)" unfolding jump_def jumpF_def using that filterlim_at_top_at_bot[OF _ _ trivial_limit_at_right_real] by auto moreover have ?thesis when "\ filterlim f at_bot (at_left a) \ \ filterlim f at_top (at_right a)" "\ filterlim f at_top (at_left a) \ \ filterlim f at_bot (at_right a)" proof (cases "f a=0") case False have "jumpF f (at_right a) = 0" "jumpF f (at_left a) = 0" proof - have "isCont (inverse o inverse o f) a" using cont False unfolding comp_def by (rule_tac continuous_at_within_inverse, auto) then have "isCont f a" unfolding comp_def by auto then have "(f \ f a) (at_right a)" "(f \ f a) (at_left a)" unfolding continuous_at_split by (auto simp add:continuous_within) moreover note trivial_limit_at_left_real trivial_limit_at_right_real ultimately show "jumpF f (at_right a) = 0" "jumpF f (at_left a) = 0" unfolding jumpF_def using filterlim_at_bot_nhds filterlim_at_top_nhds by metis+ qed then show ?thesis unfolding jump_def using that by auto next case True then have tends0:"((\x. inverse (f x)) \ 0) (at a)" using cont unfolding isCont_def comp_def by auto have "jump f a = 0" using that unfolding jump_def by auto have r_lim:"if r>0 then filterlim f at_top (at_right a) else filterlim f at_bot (at_right a)" proof (cases "r>0") case True then have "\\<^sub>F x in (at_right a). 0 < f x" using sgnxr unfolding has_sgnx_def by (auto elim:eventually_mono) then have "filterlim f at_top (at_right a)" using filterlim_inverse_at_top[of "\x. inverse (f x)", simplified] tends0 unfolding filterlim_at_split by auto then show ?thesis using True by presburger next case False then have "\\<^sub>F x in (at_right a). 0 > f x" using sgnxr \r\0\ False unfolding has_sgnx_def apply (elim eventually_mono) by (meson linorder_neqE_linordered_idom sgn_less) then have "filterlim f at_bot (at_right a)" using filterlim_inverse_at_bot[of "\x. inverse (f x)", simplified] tends0 unfolding filterlim_at_split by auto then show ?thesis using False by simp qed have l_lim:"if l>0 then filterlim f at_top (at_left a) else filterlim f at_bot (at_left a)" proof (cases "l>0") case True then have "\\<^sub>F x in (at_left a). 0 < f x" using sgnxl unfolding has_sgnx_def by (auto elim:eventually_mono) then have "filterlim f at_top (at_left a)" using filterlim_inverse_at_top[of "\x. inverse (f x)", simplified] tends0 unfolding filterlim_at_split by auto then show ?thesis using True by presburger next case False then have "\\<^sub>F x in (at_left a). 0 > f x" using sgnxl \l\0\ False unfolding has_sgnx_def apply (elim eventually_mono) by (meson linorder_neqE_linordered_idom sgn_less) then have "filterlim f at_bot (at_left a)" using filterlim_inverse_at_bot[of "\x. inverse (f x)", simplified] tends0 unfolding filterlim_at_split by auto then show ?thesis using False by simp qed have ?thesis when "l>0" "r>0" using that l_lim r_lim \jump f a=0\ unfolding jumpF_def by auto moreover have ?thesis when "\ l>0" "\ r>0" proof - have "filterlim f at_bot (at_right a)" "filterlim f at_bot (at_left a)" using r_lim l_lim that by auto moreover then have "\ filterlim f at_top (at_right a)" "\ filterlim f at_top (at_left a)" by (auto elim: filterlim_at_top_at_bot) ultimately have "jumpF f (at_right a) = -1/2 " "jumpF f (at_left a) = -1/2" unfolding jumpF_def by auto then show ?thesis using \jump f a=0\ by auto qed moreover have ?thesis when "l>0" "\ r>0" proof - note \\ filterlim f at_top (at_left a) \ \ filterlim f at_bot (at_right a)\ moreover have "filterlim f at_bot (at_right a)" "filterlim f at_top (at_left a)" using r_lim l_lim that by auto ultimately have False by auto then show ?thesis by auto qed moreover have ?thesis when "\ l>0" "r>0" proof - note \\ filterlim f at_bot (at_left a) \ \ filterlim f at_top (at_right a)\ moreover have "filterlim f at_bot (at_left a)" "filterlim f at_top (at_right a)" using r_lim l_lim that by auto ultimately have False by auto then show ?thesis by auto qed ultimately show ?thesis by auto qed ultimately show ?thesis by auto qed lemma jump_linear_comp: assumes "c\0" shows "jump (f o (\x. c*x+b)) x = (if c>0 then jump f (c*x+b) else -jump f (c*x+b))" proof (cases "c>0") case False then have "c<0" using \c\0\ by auto let ?g = "\x. c*x+b" have "filtermap ?g (at_left x) = at_right (?g x)" "filtermap ?g (at_right x) = at_left (?g x)" using \c<0\ filtermap_linear_at_left[OF \c\0\, of b x] filtermap_linear_at_right[OF \c\0\, of b x] by auto then have "jump (f \ ?g) x = - jump f (c * x + b)" unfolding jump_def filterlim_def comp_def apply (auto simp add: filtermap_filtermap[of f ?g,symmetric]) apply (fold filterlim_def) by (auto elim:filterlim_at_top_at_bot) then show ?thesis using \c<0\ by auto next case True let ?g = "\x. c*x+b" have "filtermap ?g (at_left x) = at_left (?g x)" "filtermap ?g (at_right x) = at_right (?g x)" using True filtermap_linear_at_left[OF \c\0\, of b x] filtermap_linear_at_right[OF \c\0\, of b x] by auto then have "jump (f \ ?g) x = jump f (c * x + b)" unfolding jump_def filterlim_def comp_def by (auto simp add: filtermap_filtermap[of f ?g,symmetric]) then show ?thesis using True by auto qed lemma jump_divide_derivative: assumes "isCont f x" "g x = 0" "f x\0" and g_deriv:"(g has_field_derivative c) (at x)" and "c\0" shows "jump (\t. f t/g t) x = (if sgn c = sgn ( f x) then 1 else -1)" proof - have g_tendsto:"(g \ 0) (at_left x)" "(g \ 0) (at_right x)" by (metis DERIV_isCont Lim_at_imp_Lim_at_within assms(2) assms(4) continuous_at)+ have f_tendsto:"(f \ f x) (at_left x)" "(f \ f x) (at_right x)" using Lim_at_imp_Lim_at_within assms(1) continuous_at by blast+ have ?thesis when "c>0" "f x>0" proof - have "(g has_sgnx - sgn (f x)) (at_left x)" using has_sgnx_derivative_at_left[OF g_deriv \g x=0\] that by auto moreover have "(g has_sgnx sgn (f x)) (at_right x)" using has_sgnx_derivative_at_right[OF g_deriv \g x=0\] that by auto ultimately have "(LIM t at_left x. f t / g t :> at_bot) \ (LIM t at_right x. f t / g t :> at_top)" using filterlim_divide_at_bot_at_top_iff[OF _ \f x\0\, of f] using f_tendsto(1) f_tendsto(2) g_tendsto(1) g_tendsto(2) by blast moreover have "sgn c = sgn (f x)" using that by auto ultimately show ?thesis unfolding jump_def by auto qed moreover have ?thesis when "c>0" "f x<0" proof - have "(g has_sgnx sgn (f x)) (at_left x)" using has_sgnx_derivative_at_left[OF g_deriv \g x=0\] that by auto moreover have "(g has_sgnx - sgn (f x)) (at_right x)" using has_sgnx_derivative_at_right[OF g_deriv \g x=0\] that by auto ultimately have "(LIM t at_left x. f t / g t :> at_top) \ (LIM t at_right x. f t / g t :> at_bot)" using filterlim_divide_at_bot_at_top_iff[OF _ \f x\0\, of f] using f_tendsto(1) f_tendsto(2) g_tendsto(1) g_tendsto(2) by blast moreover from this have "\ (LIM t at_left x. f t / g t :> at_bot)" using filterlim_at_top_at_bot by fastforce moreover have "sgn c \ sgn (f x)" using that by auto ultimately show ?thesis unfolding jump_def by auto qed moreover have ?thesis when "c<0" "f x>0" proof - have "(g has_sgnx sgn (f x)) (at_left x)" using has_sgnx_derivative_at_left[OF g_deriv \g x=0\] that by auto moreover have "(g has_sgnx - sgn (f x)) (at_right x)" using has_sgnx_derivative_at_right[OF g_deriv \g x=0\] that by auto ultimately have "(LIM t at_left x. f t / g t :> at_top) \ (LIM t at_right x. f t / g t :> at_bot)" using filterlim_divide_at_bot_at_top_iff[OF _ \f x\0\, of f] using f_tendsto(1) f_tendsto(2) g_tendsto(1) g_tendsto(2) by blast moreover from this have "\ (LIM t at_left x. f t / g t :> at_bot)" using filterlim_at_top_at_bot by fastforce moreover have "sgn c \ sgn (f x)" using that by auto ultimately show ?thesis unfolding jump_def by auto qed moreover have ?thesis when "c<0" "f x<0" proof - have "(g has_sgnx - sgn (f x)) (at_left x)" using has_sgnx_derivative_at_left[OF g_deriv \g x=0\] that by auto moreover have "(g has_sgnx sgn (f x)) (at_right x)" using has_sgnx_derivative_at_right[OF g_deriv \g x=0\] that by auto ultimately have "(LIM t at_left x. f t / g t :> at_bot) \ (LIM t at_right x. f t / g t :> at_top)" using filterlim_divide_at_bot_at_top_iff[OF _ \f x\0\, of f] using f_tendsto(1) f_tendsto(2) g_tendsto(1) g_tendsto(2) by blast moreover have "sgn c =sgn (f x)" using that by auto ultimately show ?thesis unfolding jump_def by auto qed ultimately show ?thesis using \c\0\ \f x\0\ by argo qed lemma jump_jump_poly: "jump (\x. poly q x / poly p x) a = jump_poly q p a" proof (cases "p=0") case True then show ?thesis by auto next case False obtain p' q' where p':"p= p'*gcd p q" and q':"q=q'*gcd p q" using gcd_dvd1 gcd_dvd2 dvd_def[of "gcd p q", simplified mult.commute] by metis then have "coprime p' q'" "p'\0" "gcd p q\0" using gcd_coprime \p\0\ by auto define f where "f \ (\x. poly q' x / poly p' x)" define g where "g \ (\x. if poly (gcd p q) x = 0 then 0::real else 1)" have g_tendsto:"(g \ 1) (at_left a)" "(g \ 1) (at_right a)" proof - have "(poly (gcd p q) has_sgnx 1) (at_left a) \ (poly (gcd p q) has_sgnx - 1) (at_left a)" "(poly (gcd p q) has_sgnx 1) (at_right a) \ (poly (gcd p q) has_sgnx - 1) (at_right a)" using \p\0\ poly_has_sgnx_values by auto then have " \\<^sub>F x in at_left a. g x = 1" " \\<^sub>F x in at_right a. g x = 1" unfolding has_sgnx_def g_def by (auto elim:eventually_mono) then show "(g \ 1) (at_left a)" "(g \ 1) (at_right a)" using tendsto_eventually by auto qed have "poly q x / poly p x = g x * f x" for x unfolding f_def g_def by (subst p',subst q',auto) then have "jump (\x. poly q x / poly p x) a = jump (\x. g x * f x) a" by auto also have "... = jump f a" unfolding jump_def apply (subst (1 2) filterlim_tendsto_pos_mult_at_top_iff) prefer 5 apply (subst (1 2) filterlim_tendsto_pos_mult_at_bot_iff) using g_tendsto by auto also have "... = jump_poly q' p' a" using jump_jump_poly_aux[OF \p'\0\ \coprime p' q'\] unfolding f_def by auto also have "... = jump_poly q p a" using jump_poly_mult[OF \gcd p q \ 0\, of q'] p' q' by (metis mult.commute) finally show ?thesis . qed lemma jump_Im_divide_Re_0: assumes "path g" "Re (g x)\0" "0t. Im (g t) / Re (g t)) x = 0" proof - have "isCont g x" using \path g\[unfolded path_def] \0 \x<1\ apply (elim continuous_on_interior) by auto then have "isCont (\t. Im(g t)/Re(g t)) x" using \Re (g x)\0\ by (auto intro:continuous_intros isCont_Re isCont_Im) then show "jump (\t. Im(g t)/Re(g t)) x=0" using jump_not_infinity by auto qed lemma jumpF_im_divide_Re_0: assumes "path g" "Re (g x)\0" shows "\0\x;x<1\ \ jumpF (\t. Im (g t) / Re (g t)) (at_right x) = 0" "\01\ \ jumpF (\t. Im (g t) / Re (g t)) (at_left x) = 0" proof - define g' where "g' = (\t. Im (g t) / Re (g t))" show "jumpF g' (at_right x) = 0" when "0\x" "x<1" proof - have "(g' \ g' x) (at_right x)" proof (cases "x=0") case True have "continuous (at_right 0) g" using \path g\ unfolding path_def by (auto elim:continuous_on_at_right) then have "continuous (at_right x) (\t. Im(g t))" "continuous (at_right x) (\t. Re(g t))" using continuous_Im continuous_Re True by auto moreover have "Re (g (netlimit (at_right x))) \ 0" using assms(2) by (simp add: Lim_ident_at) ultimately have "continuous (at_right x) (\t. Im (g t)/Re(g t))" by (auto intro:continuous_divide) then show ?thesis unfolding g'_def continuous_def by (simp add: Lim_ident_at) next case False have "isCont (\x. Im (g x)) x" "isCont (\x. Re (g x)) x" using \path g\ unfolding path_def by (metis False atLeastAtMost_iff at_within_Icc_at continuous_Im continuous_Re continuous_on_eq_continuous_within less_le that)+ then have "isCont g' x" using assms(2) unfolding g'_def by (auto intro:continuous_intros) then show ?thesis unfolding isCont_def using filterlim_at_split by blast qed then have "\ filterlim g' at_top (at_right x)" "\ filterlim g' at_bot (at_right x)" using filterlim_at_top_nhds[of g' "at_right x"] filterlim_at_bot_nhds[of g' "at_right x"] by auto then show ?thesis unfolding jumpF_def by auto qed show "jumpF g' (at_left x) = 0" when "01" proof - have "(g' \ g' x) (at_left x)" proof (cases "x=1") case True have "continuous (at_left 1) g" using \path g\ unfolding path_def by (auto elim:continuous_on_at_left) then have "continuous (at_left x) (\t. Im(g t))" "continuous (at_left x) (\t. Re(g t))" using continuous_Im continuous_Re True by auto moreover have "Re (g (netlimit (at_left x))) \ 0" using assms(2) by (simp add: Lim_ident_at) ultimately have "continuous (at_left x) (\t. Im (g t)/Re(g t))" by (auto intro:continuous_divide) then show ?thesis unfolding g'_def continuous_def by (simp add: Lim_ident_at) next case False have "isCont (\x. Im (g x)) x" "isCont (\x. Re (g x)) x" using \path g\ unfolding path_def by (metis False atLeastAtMost_iff at_within_Icc_at continuous_Im continuous_Re continuous_on_eq_continuous_within less_le that)+ then have "isCont g' x" using assms(2) unfolding g'_def by (auto) then show ?thesis unfolding isCont_def using filterlim_at_split by blast qed then have "\ filterlim g' at_top (at_left x)" "\ filterlim g' at_bot (at_left x)" using filterlim_at_top_nhds[of g' "at_left x"] filterlim_at_bot_nhds[of g' "at_left x"] by auto then show ?thesis unfolding jumpF_def by auto qed qed lemma jump_cong: assumes "x=y" and "eventually (\x. f x=g x) (at x)" shows "jump f x = jump g y" proof - have left:"eventually (\x. f x=g x) (at_left x)" and right:"eventually (\x. f x=g x) (at_right x)" using assms(2) eventually_at_split by blast+ from filterlim_cong[OF _ _ this(1)] filterlim_cong[OF _ _ this(2)] show ?thesis unfolding jump_def using assms(1) by fastforce qed lemma jumpF_cong: assumes "F=G" and "eventually (\x. f x=g x) F" shows "jumpF f F = jumpF g G" proof - have "\\<^sub>F r in G. f r = g r" using assms(1) assms(2) by force then show ?thesis by (simp add: assms(1) filterlim_cong jumpF_def) qed lemma jump_at_left_at_right_eq: assumes "isCont f x" and "f x \ 0" and sgnx_eq:"sgnx g (at_left x) = sgnx g (at_right x)" shows "jump (\t. f t/g t) x = 0" proof - define c where "c = sgn (f x)" then have "c\0" using \f x\0\ by (simp add: sgn_zero_iff) have f_tendsto:"(f \ f x) (at_left x)" " (f \ f x) (at_right x)" using \isCont f x\ Lim_at_imp_Lim_at_within isCont_def by blast+ have False when "(g has_sgnx - c) (at_left x)" "(g has_sgnx c) (at_right x)" proof - have "sgnx g (at_left x) = -c" using that(1) by auto moreover have "sgnx g (at_right x) = c" using that(2) by auto ultimately show False using sgnx_eq \c\0\ by force qed moreover have False when "(g has_sgnx c) (at_left x)" "(g has_sgnx - c) (at_right x)" proof - have "sgnx g (at_left x) = c" using that(1) by auto moreover have "sgnx g (at_right x) = - c" using that(2) by auto ultimately show False using sgnx_eq \c\0\ by force qed ultimately show ?thesis unfolding jump_def by (auto simp add:f_tendsto filterlim_divide_at_bot_at_top_iff[OF _ \f x \ 0\] c_def) qed lemma jumpF_pos_has_sgnx: assumes "jumpF f F > 0" shows "(f has_sgnx 1) F" proof - have "filterlim f at_top F" using assms unfolding jumpF_def by argo then have "eventually (\x. f x>0) F" using filterlim_at_top_dense[of f F] by blast then show ?thesis unfolding has_sgnx_def apply (elim eventually_mono) by auto qed lemma jumpF_neg_has_sgnx: assumes "jumpF f F < 0" shows "(f has_sgnx -1) F" proof - have "filterlim f at_bot F" using assms unfolding jumpF_def by argo then have "eventually (\x. f x<0) F" using filterlim_at_bot_dense[of f F] by blast then show ?thesis unfolding has_sgnx_def apply (elim eventually_mono) by auto qed lemma jumpF_IVT: fixes f::"real \ real" and a b::real defines "right\(\(R::real \ real \ bool). R (jumpF f (at_right a)) 0 \ (continuous (at_right a) f \ R (f a) 0))" and "left\(\(R::real \ real \ bool). R (jumpF f (at_left b)) 0 \ (continuous (at_left b) f \ R (f b) 0))" assumes "a left less \ right less \ left greater" shows "\x. a x f x =0" proof - have ?thesis when "right greater" "left less" proof - have "(f has_sgnx 1) (at_right a)" proof - have ?thesis when "jumpF f (at_right a)>0" using jumpF_pos_has_sgnx[OF that] . moreover have ?thesis when "f a > 0" "continuous (at_right a) f" proof - have "(f \ f a) (at_right a)" using that(2) by (simp add: continuous_within) then show ?thesis using tendsto_nonzero_has_sgnx[of f "f a" "at_right a"] that by auto qed ultimately show ?thesis using that(1) unfolding right_def by auto qed then obtain a' where "ay. a y < a' \ f y > 0" unfolding has_sgnx_def eventually_at_right using sgn_1_pos by auto have "(f has_sgnx - 1) (at_left b)" proof - have ?thesis when "jumpF f (at_left b)<0" using jumpF_neg_has_sgnx[OF that] . moreover have ?thesis when "f b < 0" "continuous (at_left b) f" proof - have "(f \ f b) (at_left b)" using that(2) by (simp add: continuous_within) then show ?thesis using tendsto_nonzero_has_sgnx[of f "f b" "at_left b"] that by auto qed ultimately show ?thesis using that(2) unfolding left_def by auto qed then obtain b' where "b'y. b' y < b \ f y < 0" unfolding has_sgnx_def eventually_at_left using sgn_1_neg by auto have "a' \ b'" proof (rule ccontr) assume "\ a' \ b'" then have "{a<.. {b'<.. {}" using \a \b' \a by auto then obtain c where "c\{a<..{b'<..0" "f c<0" using a'_def b'_def by auto then show False by auto qed define a0 where "a0=(a+a')/2" define b0 where "b0=(b+b')/2" have [simp]:"aa \b' \a'\b'\ by auto have "f a0>0" "f b0<0" using a'_def[rule_format,of a0] b'_def[rule_format,of b0] by auto moreover have "continuous_on {a0..b0} f" using cont \a < a0\ \b0 < b\ by (meson atLeastAtMost_subseteq_greaterThanLessThan_iff continuous_on_subset) ultimately have "\x>a0. x < b0 \ f x = 0" using IVT_strict[of 0 f a0 b0] by auto then show ?thesis using \a < a0\ \b0 < b\ by (meson lessThan_strict_subset_iff psubsetE subset_psubset_trans) qed moreover have ?thesis when "right less" "left greater" proof - have "(f has_sgnx -1) (at_right a)" proof - have ?thesis when "jumpF f (at_right a)<0" using jumpF_neg_has_sgnx[OF that] . moreover have ?thesis when "f a < 0" "continuous (at_right a) f" proof - have "(f \ f a) (at_right a)" using that(2) by (simp add: continuous_within) then show ?thesis using tendsto_nonzero_has_sgnx[of f "f a" "at_right a"] that by auto qed ultimately show ?thesis using that(1) unfolding right_def by auto qed then obtain a' where "ay. a y < a' \ f y < 0" unfolding has_sgnx_def eventually_at_right using sgn_1_neg by auto have "(f has_sgnx 1) (at_left b)" proof - have ?thesis when "jumpF f (at_left b)>0" using jumpF_pos_has_sgnx[OF that] . moreover have ?thesis when "f b > 0" "continuous (at_left b) f" proof - have "(f \ f b) (at_left b)" using that(2) by (simp add: continuous_within) then show ?thesis using tendsto_nonzero_has_sgnx[of f "f b" "at_left b"] that by auto qed ultimately show ?thesis using that(2) unfolding left_def by auto qed then obtain b' where "b'y. b' y < b \ f y > 0" unfolding has_sgnx_def eventually_at_left using sgn_1_pos by auto have "a' \ b'" proof (rule ccontr) assume "\ a' \ b'" then have "{a<.. {b'<.. {}" using \a \b' \a by auto then obtain c where "c\{a<..{b'<..0" "f c<0" using a'_def b'_def by auto then show False by auto qed define a0 where "a0=(a+a')/2" define b0 where "b0=(b+b')/2" have [simp]:"aa \b' \a'\b'\ by auto have "f a0<0" "f b0>0" using a'_def[rule_format,of a0] b'_def[rule_format,of b0] by auto moreover have "continuous_on {a0..b0} f" using cont \a < a0\ \b0 < b\ by (meson atLeastAtMost_subseteq_greaterThanLessThan_iff continuous_on_subset) ultimately have "\x>a0. x < b0 \ f x = 0" using IVT_strict[of 0 f a0 b0] by auto then show ?thesis using \a < a0\ \b0 < b\ by (meson lessThan_strict_subset_iff psubsetE subset_psubset_trans) qed ultimately show ?thesis using right_left by auto qed lemma jumpF_eventually_const: assumes "eventually (\x. f x=c) F" "F\bot" shows "jumpF f F = 0" proof - have "jumpF f F = jumpF (\_. c) F" apply (rule jumpF_cong) using assms(1) by auto also have "... = 0" using jumpF_const[OF \F\bot\] by simp finally show ?thesis . qed lemma jumpF_tan_comp: "jumpF (f o tan) (at_right x) = (if cos x = 0 then jumpF f at_bot else jumpF f (at_right (tan x)))" "jumpF (f o tan) (at_left x) = (if cos x =0 then jumpF f at_top else jumpF f (at_left (tan x)))" proof - have "filtermap (f \ tan) (at_right x) = (if cos x = 0 then filtermap f at_bot else filtermap f (at_right (tan x)))" unfolding comp_def apply (subst filtermap_filtermap[of f tan,symmetric]) using filtermap_tan_at_right_inf filtermap_tan_at_right by auto then show "jumpF (f o tan) (at_right x) = (if cos x = 0 then jumpF f at_bot else jumpF f (at_right (tan x)))" unfolding jumpF_def filterlim_def by auto next have "filtermap (f \ tan) (at_left x) = (if cos x = 0 then filtermap f at_top else filtermap f (at_left (tan x)))" unfolding comp_def apply (subst filtermap_filtermap[of f tan,symmetric]) using filtermap_tan_at_left_inf filtermap_tan_at_left by auto then show "jumpF (f o tan) (at_left x) = (if cos x = 0 then jumpF f at_top else jumpF f (at_left (tan x)))" unfolding jumpF_def filterlim_def by auto qed subsection \Finite jumpFs over an interval\ definition finite_jumpFs::"(real \ real) \ real \ real \ bool" where "finite_jumpFs f a b = finite {x. (jumpF f (at_left x) \0 \ jumpF f (at_right x) \0) \ a\x \ x\b}" lemma finite_jumpFs_linear_pos: assumes "c>0" shows "finite_jumpFs (f o (\x. c * x + b)) lb ub \ finite_jumpFs f (c * lb +b) (c * ub + b)" proof - define left where "left = (\f lb ub. {x. jumpF f (at_left x) \ 0 \ lb \ x \ x \ ub})" define right where "right = (\f lb ub. {x. jumpF f (at_right x) \ 0 \ lb \ x \ x \ ub})" define g where "g=(\x. c*x+b)" define gi where "gi = (\x. (x-b)/c)" have "finite_jumpFs (f o (\x. c * x + b)) lb ub = finite (left (f o g) lb ub \ right (f o g) lb ub)" unfolding finite_jumpFs_def apply (rule arg_cong[where f=finite]) by (auto simp add:left_def right_def g_def) also have "... = finite (gi ` (left f (g lb) (g ub) \ right f (g lb) (g ub)))" proof - have j_rw: "jumpF (f o g) (at_left x) = jumpF f (at_left (g x))" "jumpF (f o g) (at_right x) = jumpF f (at_right (g x))" for x using jumpF_linear_comp[of c f b x] \c>0\ unfolding g_def by auto then have "left (f o g) lb ub = gi ` left f (g lb) (g ub)" "right (f o g) lb ub = gi ` right f (g lb) (g ub)" unfolding left_def right_def gi_def using \c>0\ by (auto simp add:g_def field_simps) then have "left (f o g) lb ub \ right (f o g) lb ub = gi ` (left f (g lb) (g ub) \ right f (g lb) (g ub))" by auto then show ?thesis by auto qed also have "... = finite (left f (g lb) (g ub) \ right f (g lb) (g ub))" apply (rule finite_image_iff) unfolding gi_def using \c >0\ inj_on_def by fastforce also have "... = finite_jumpFs f (c * lb +b) (c * ub + b)" unfolding finite_jumpFs_def apply (rule arg_cong[where f=finite]) by (auto simp add:left_def right_def g_def) finally show ?thesis . qed lemma finite_jumpFs_consts: "finite_jumpFs (\_ . c) lb ub" unfolding finite_jumpFs_def using jumpF_const by auto lemma finite_jumpFs_combine: assumes "finite_jumpFs f a b" "finite_jumpFs f b c" shows "finite_jumpFs f a c" proof - define P where "P=(\x. jumpF f (at_left x) \ 0 \ jumpF f (at_right x) \ 0)" have "{x. P x \ a \ x \ x \ c} \ {x. P x \ a \ x \ x\b} \ {x. P x \ b \x \ x\c}" by auto moreover have "finite ({x. P x \ a \ x \ x\b} \ {x. P x \ b \x \ x\c})" using assms unfolding finite_jumpFs_def P_def by auto ultimately have "finite {x. P x \ a \ x \ x \ c}" using finite_subset by auto then show ?thesis unfolding finite_jumpFs_def P_def by auto qed lemma finite_jumpFs_subE: assumes "finite_jumpFs f a b" "a\a'" "b'\b" shows "finite_jumpFs f a' b'" using assms unfolding finite_jumpFs_def apply (elim rev_finite_subset) by auto lemma finite_Psegments_Re_imp_jumpFs: assumes "finite_Psegments (\t. Re (g t - z) = 0) a b" "continuous_on {a..b} g" shows "finite_jumpFs (\t. Im (g t - z)/Re (g t - z)) a b" using assms proof (induct rule:finite_Psegments.induct) case (emptyI a b) then show ?case unfolding finite_jumpFs_def by (auto intro:rev_finite_subset[of "{a}"]) next case (insertI_1 s a b) define f where "f=(\t. Im (g t - z) / Re (g t - z))" have "finite_jumpFs f a s" proof - have "continuous_on {a..s} g" using \continuous_on {a..b} g\ \s \ {a.. by (auto elim:continuous_on_subset) then show ?thesis using insertI_1 unfolding f_def by auto qed moreover have "finite_jumpFs f s b" proof - have "jumpF f (at_left x) =0" "jumpF f (at_right x) = 0" when "x\{s<.. 0 \ jumpF f (at_right x) \ 0) \ s \ x \ x \ b} = {x. (jumpF f (at_left x) \ 0 \ jumpF f (at_right x) \ 0) \ (x=s \ x = b)}" using \s\{a.. by force then show ?thesis unfolding finite_jumpFs_def by auto qed ultimately show ?case using finite_jumpFs_combine[of _ a s b] unfolding f_def by auto next case (insertI_2 s a b) define f where "f=(\t. Im (g t - z) / Re (g t - z))" have "finite_jumpFs f a s" proof - have "continuous_on {a..s} g" using \continuous_on {a..b} g\ \s \ {a.. by (auto elim:continuous_on_subset) then show ?thesis using insertI_2 unfolding f_def by auto qed moreover have "finite_jumpFs f s b" proof - have "jumpF f (at_left x) =0" "jumpF f (at_right x) = 0" when "x\{s<..continuous_on {a..b} g\]) using insertI_2.hyps(1) that apply auto[2] using insertI_2.hyps(3) that by blast then show "jumpF f (at_left x) =0" "jumpF f (at_right x) = 0" by (simp_all add: continuous_at_split jumpF_not_infinity) qed then have "{x. (jumpF f (at_left x) \ 0 \ jumpF f (at_right x) \ 0) \ s \ x \ x \ b} = {x. (jumpF f (at_left x) \ 0 \ jumpF f (at_right x) \ 0) \ (x=s \ x = b)}" using \s\{a.. by force then show ?thesis unfolding finite_jumpFs_def by auto qed ultimately show ?case using finite_jumpFs_combine[of _ a s b] unfolding f_def by auto qed lemma finite_ReZ_segments_imp_jumpFs: assumes "finite_ReZ_segments g z" "path g" shows "finite_jumpFs (\t. Im (g t - z)/Re (g t - z)) 0 1" using assms unfolding finite_ReZ_segments_def path_def by (rule finite_Psegments_Re_imp_jumpFs) subsection \@{term jumpF} at path ends\ definition jumpF_pathstart::"(real \ complex) \ complex \ real" where "jumpF_pathstart g z= jumpF (\t. Im(g t- z)/Re(g t - z)) (at_right 0)" definition jumpF_pathfinish::"(real \ complex) \ complex \ real" where "jumpF_pathfinish g z= jumpF (\t. Im(g t - z)/Re(g t -z)) (at_left 1)" lemma jumpF_pathstart_eq_0: assumes "path g" "Re(pathstart g)\Re z" shows "jumpF_pathstart g z = 0" unfolding jumpF_pathstart_def apply (rule jumpF_im_divide_Re_0) using assms[unfolded pathstart_def] by auto lemma jumpF_pathfinish_eq_0: assumes "path g" "Re(pathfinish g)\Re z" shows "jumpF_pathfinish g z = 0" unfolding jumpF_pathfinish_def apply (rule jumpF_im_divide_Re_0) using assms[unfolded pathfinish_def] by auto lemma shows jumpF_pathfinish_reversepath: "jumpF_pathfinish (reversepath g) z = jumpF_pathstart g z" and jumpF_pathstart_reversepath: "jumpF_pathstart (reversepath g) z = jumpF_pathfinish g z" proof - define f where "f=(\t. Im (g t - z) / Re (g t - z))" define f' where "f'=(\t. Im (reversepath g t - z) / Re (reversepath g t - z))" have "f o (\t. 1 - t) = f'" unfolding f_def f'_def comp_def reversepath_def by auto then show "jumpF_pathfinish (reversepath g) z = jumpF_pathstart g z" "jumpF_pathstart (reversepath g) z = jumpF_pathfinish g z" unfolding jumpF_pathstart_def jumpF_pathfinish_def using jumpF_linear_comp(2)[of "-1" f 1 0,simplified] jumpF_linear_comp(1)[of "-1" f 1 1,simplified] apply (fold f_def f'_def) by auto qed lemma jumpF_pathstart_joinpaths[simp]: "jumpF_pathstart (g1+++g2) z = jumpF_pathstart g1 z" proof - let ?h="(\t. Im (g1 t - z) / Re (g1 t - z))" let ?f="\t. Im ((g1 +++ g2) t - z) / Re ((g1 +++ g2) t - z)" have "jumpF_pathstart g1 z = jumpF ?h (at_right 0)" unfolding jumpF_pathstart_def by simp also have "... = jumpF (?h o (\t. 2*t)) (at_right 0)" using jumpF_linear_comp[of 2 ?h 0 0,simplified] by auto also have "... = jumpF ?f (at_right 0)" proof (rule jumpF_cong) show " \\<^sub>F x in at_right 0. (?h \ (*) 2) x =?f x" unfolding eventually_at_right apply (intro exI[where x="1/2"]) by (auto simp add:joinpaths_def) qed simp also have "... =jumpF_pathstart (g1+++g2) z" unfolding jumpF_pathstart_def by simp finally show ?thesis by simp qed lemma jumpF_pathfinish_joinpaths[simp]: "jumpF_pathfinish (g1+++g2) z = jumpF_pathfinish g2 z" proof - let ?h="(\t. Im (g2 t - z) / Re (g2 t - z))" let ?f="\t. Im ((g1 +++ g2) t - z) / Re ((g1 +++ g2) t - z)" have "jumpF_pathfinish g2 z = jumpF ?h (at_left 1)" unfolding jumpF_pathfinish_def by simp also have "... = jumpF (?h o (\t. 2*t-1)) (at_left 1)" using jumpF_linear_comp[of 2 _ "-1" 1,simplified] by auto also have "... = jumpF ?f (at_left 1)" proof (rule jumpF_cong) show " \\<^sub>F x in at_left 1. (?h \ (\t. 2 * t - 1)) x =?f x" unfolding eventually_at_left apply (intro exI[where x="1/2"]) by (auto simp add:joinpaths_def) qed simp also have "... =jumpF_pathfinish (g1+++g2) z" unfolding jumpF_pathfinish_def by simp finally show ?thesis by simp qed subsection \Cauchy index\ \\Deprecated, use "cindexE" if possible\ definition cindex::"real \ real \ (real \ real) \ int" where "cindex a b f = (\x\{x. jump f x\0 \ a x real \ (real \ real) \ real" where "cindexE a b f = (\x\{x. jumpF f (at_right x) \0 \ a\x \ xx\{x. jumpF f (at_left x) \0 \ a x\b}. jumpF f (at_left x))" definition cindexE_ubd::"(real \ real) \ real" where "cindexE_ubd f = (\x\{x. jumpF f (at_right x) \0 }. jumpF f (at_right x)) - (\x\{x. jumpF f (at_left x) \0}. jumpF f (at_left x))" lemma cindexE_empty: "cindexE a a f = 0" unfolding cindexE_def by (simp add: sum.neutral) lemma cindex_const: "cindex a b (\_. c) = 0" unfolding cindex_def apply (rule sum.neutral) by auto lemma cindex_eq_cindex_poly: "cindex a b (\x. poly q x/poly p x) = cindex_poly a b q p" proof (cases "p=0") case True then show ?thesis using cindex_const by auto next case False have "cindex_poly a b q p = (\x |jump_poly q p x \0 \ a < x \ x < b. jump_poly q p x)" unfolding cindex_poly_def apply (rule sum.mono_neutral_cong_right) using jump_poly_not_root by (auto simp add: \p\0\ poly_roots_finite) also have "... = cindex a b (\x. poly q x/poly p x)" unfolding cindex_def apply (rule sum.cong) using jump_jump_poly[of q] by auto finally show ?thesis by auto qed lemma cindex_combine: assumes finite:"finite {x. jump f x\0 \ a xs. sum (jump f) ({x. jump f x\0 \ a x s))" have ssum_union:"ssum (A \ B) = ssum A + ssum B" when "A \ B ={}" for A B proof - define C where "C={x. jump f x \ 0 \ a x A" "C \ B"] that by (simp add: inf_assoc inf_sup_aci(3) inf_sup_distrib1 sum.union_disjoint) qed have "cindex a c f = ssum ({a<.. {b} \ {b<..a \b by fastforce moreover have "cindex a b f = ssum {a<..a \b by (intro sum.cong,auto) moreover have "jump f b = ssum {b}" unfolding ssum_def using \a \b by (cases "jump f b=0",auto) moreover have "cindex b c f = ssum {b<..a \b by (intro sum.cong,auto) ultimately show ?thesis apply (subst (asm) ssum_union,simp) by (subst (asm) ssum_union,auto) qed lemma cindexE_combine: assumes finite:"finite_jumpFs f a c" and "a\b" "b\c" shows "cindexE a c f = cindexE a b f + cindexE b c f" proof - define S where "S={x. (jumpF f (at_left x) \ 0 \ jumpF f (at_right x) \ 0) \ a \ x \ x \ c}" define A0 where "A0={x. jumpF f (at_right x) \ 0 \ a \ x \ x < c}" define A1 where "A1={x. jumpF f (at_right x) \ 0 \ a \ x \ x < b}" define A2 where "A2={x. jumpF f (at_right x) \ 0 \ b \ x \ x < c}" define B0 where "B0={x. jumpF f (at_left x) \ 0 \ a < x \ x \ c}" define B1 where "B1={x. jumpF f (at_left x) \ 0 \ a < x \ x \ b}" define B2 where "B2={x. jumpF f (at_left x) \ 0 \ b < x \ x \ c}" have [simp]:"finite A1" "finite A2" "finite B1" "finite B2" proof - have "finite S" using finite unfolding finite_jumpFs_def S_def by auto moreover have "A1 \ S" "A2 \ S" "B1 \ S" "B2 \ S" unfolding A1_def A2_def B1_def B2_def S_def using \a\b\ \b\c\ by auto ultimately show "finite A1" "finite A2" "finite B1" "finite B2" by (auto elim:finite_subset) qed have "cindexE a c f = sum (\x. jumpF f (at_right x)) A0 - sum (\x. jumpF f (at_left x)) B0" unfolding cindexE_def A0_def B0_def by auto also have "... = sum (\x. jumpF f (at_right x)) (A1 \ A2) - sum (\x. jumpF f (at_left x)) (B1 \ B2)" proof - have "A0=A1\A2" unfolding A0_def A1_def A2_def using assms by auto moreover have "B0=B1\B2" unfolding B0_def B1_def B2_def using assms by auto ultimately show ?thesis by auto qed also have "... = cindexE a b f + cindexE b c f" proof - have "A1 \ A2 = {}" unfolding A1_def A2_def by auto moreover have "B1 \ B2 = {}" unfolding B1_def B2_def by auto ultimately show ?thesis unfolding cindexE_def apply (fold A1_def A2_def B1_def B2_def) by (auto simp add:sum.union_disjoint) qed finally show ?thesis . qed lemma cindex_linear_comp: assumes "c\0" shows "cindex lb ub (f o (\x. c*x+b)) = (if c>0 then cindex (c*lb+b) (c*ub+b) f else - cindex (c*ub+b) (c*lb+b) f)" proof (cases "c>0") case False then have "c<0" using \c\0\ by auto have "cindex lb ub (f o (\x. c*x+b)) = - cindex (c*ub+b) (c*lb+b) f" unfolding cindex_def apply (subst sum_negf[symmetric]) apply (intro sum.reindex_cong[of "\x. (x-b)/c"]) subgoal by (simp add: inj_on_def) subgoal using False apply (subst jump_linear_comp[OF \c\0\]) by (auto simp add:\c<0\ \c\0\ field_simps) subgoal for x apply (subst jump_linear_comp[OF \c\0\]) by (auto simp add:\c<0\ \c\0\ False field_simps) done then show ?thesis using False by auto next case True have "cindex lb ub (f o (\x. c*x+b)) = cindex (c*lb+b) (c*ub+b) f" unfolding cindex_def apply (intro sum.reindex_cong[of "\x. (x-b)/c"]) subgoal by (simp add: inj_on_def) subgoal apply (subst jump_linear_comp[OF \c\0\]) by (auto simp add: True \c\0\ field_simps) subgoal for x apply (subst jump_linear_comp[OF \c\0\]) by (auto simp add: \c\0\ True field_simps) done then show ?thesis using True by auto qed lemma cindexE_linear_comp: assumes "c\0" shows "cindexE lb ub (f o (\x. c*x+b)) = (if c>0 then cindexE (c*lb+b) (c*ub+b) f else - cindexE (c*ub+b) (c*lb+b) f)" proof - define cright where "cright = (\lb ub f. (\x | jumpF f (at_right x) \ 0 \ lb \ x \ x < ub. jumpF f (at_right x)))" define cleft where "cleft = (\lb ub f. (\x | jumpF f (at_left x) \ 0 \ lb < x \ x \ ub. jumpF f (at_left x)))" have cindexE_unfold:"cindexE lb ub f = cright lb ub f - cleft lb ub f" for lb ub f unfolding cindexE_def cright_def cleft_def by auto have ?thesis when "c<0" proof - have "cright lb ub (f \ (\x. c * x + b)) = cleft (c * ub + b) (c * lb + b) f" unfolding cright_def cleft_def apply (intro sum.reindex_cong[of "\x. (x-b)/c"]) subgoal by (simp add: inj_on_def) subgoal using that by (subst jumpF_linear_comp[OF \c\0\],auto simp add:field_simps) subgoal for x using that by (subst jumpF_linear_comp[OF \c\0\],auto simp add: field_simps) done moreover have "cleft lb ub (f \ (\x. c * x + b)) = cright (c*ub+b) (c*lb + b) f" unfolding cright_def cleft_def apply (intro sum.reindex_cong[of "\x. (x-b)/c"]) subgoal by (simp add: inj_on_def) subgoal using that by (subst jumpF_linear_comp[OF \c\0\],auto simp add:field_simps) subgoal for x using that by (subst jumpF_linear_comp[OF \c\0\],auto simp add: field_simps) done ultimately show ?thesis unfolding cindexE_unfold using that by auto qed moreover have ?thesis when "c>0" proof - have "cright lb ub (f \ (\x. c * x + b)) = cright (c * lb + b) (c * ub + b) f" unfolding cright_def cleft_def apply (intro sum.reindex_cong[of "\x. (x-b)/c"]) subgoal by (simp add: inj_on_def) subgoal using that by (subst jumpF_linear_comp[OF \c\0\],auto simp add:field_simps) subgoal for x using that by (subst jumpF_linear_comp[OF \c\0\],auto simp add: field_simps) done moreover have "cleft lb ub (f \ (\x. c * x + b)) = cleft (c*lb+b) (c*ub + b) f" unfolding cright_def cleft_def apply (intro sum.reindex_cong[of "\x. (x-b)/c"]) subgoal by (simp add: inj_on_def) subgoal using that by (subst jumpF_linear_comp[OF \c\0\],auto simp add:field_simps) subgoal for x using that by (subst jumpF_linear_comp[OF \c\0\],auto simp add: field_simps) done ultimately show ?thesis unfolding cindexE_unfold using that by auto qed ultimately show ?thesis using \c\0\ by auto qed lemma cindexE_cong: assumes "finite s" and fg_eq:"\x. \as\ \ f x = g x" shows "cindexE a b f = cindexE a b g" proof - define left where "left=(\f. (\x | jumpF f (at_left x) \ 0 \ a < x \ x \ b. jumpF f (at_left x)))" define right where "right=(\f. (\x | jumpF f (at_right x) \ 0 \ a \ x \ x < b. jumpF f (at_right x)))" have "left f = left g" proof - have "jumpF f (at_left x) = jumpF g (at_left x)" when "ab" for x proof (rule jumpF_cong) define cs where "cs \ {y\s. a y (if cs = {} then (x+a)/2 else Max cs)" have "finite cs" unfolding cs_def using assms(1) by auto have "c (\y. c y f y=g y)" proof (cases "cs={}") case True then have "\y. c y y \ s" unfolding cs_def c_def by force moreover have "c=(x+a)/2" using True unfolding c_def by auto ultimately show ?thesis using fg_eq using that by auto next case False then have "c\cs" unfolding c_def using False \finite cs\ by auto moreover have "\y. c y y \ s" proof (rule ccontr) assume "\ (\y. c < y \ y < x \ y \ s) " then obtain y' where "cs" by auto then have "y'\cs" using \c\cs\ unfolding cs_def by auto then have "y'\c" unfolding c_def using False \finite cs\ by auto then show False using \c by auto qed ultimately show ?thesis unfolding cs_def using that by (auto intro!:fg_eq) qed then show "\\<^sub>F x in at_left x. f x = g x" unfolding eventually_at_left by auto qed simp then show ?thesis unfolding left_def by (auto intro: sum.cong) qed moreover have "right f = right g" proof - have "jumpF f (at_right x) = jumpF g (at_right x)" when "a\x" "x {y\s. x y (if cs = {} then (x+b)/2 else Min cs)" have "finite cs" unfolding cs_def using assms(1) by auto have "x (\y. x y f y=g y)" proof (cases "cs={}") case True then have "\y. x y y \ s" unfolding cs_def c_def by force moreover have "c=(x+b)/2" using True unfolding c_def by auto ultimately show ?thesis using fg_eq using that by auto next case False then have "c\cs" unfolding c_def using False \finite cs\ by auto moreover have "\y. x y y \ s" proof (rule ccontr) assume "\ (\y. x < y \ y < c \ y \ s) " then obtain y' where "xs" by auto then have "y'\cs" using \c\cs\ unfolding cs_def by auto then have "y'\c" unfolding c_def using False \finite cs\ by auto then show False using \c>y'\ by auto qed ultimately show ?thesis unfolding cs_def using that by (auto intro!:fg_eq) qed then show "\\<^sub>F x in at_right x. f x = g x" unfolding eventually_at_right by auto qed simp then show ?thesis unfolding right_def by (auto intro: sum.cong) qed ultimately show ?thesis unfolding cindexE_def left_def right_def by presburger qed lemma cindexE_constI: assumes "\t. \a \ f t=c" shows "cindexE a b f = 0" proof - define left where "left=(\f. (\x | jumpF f (at_left x) \ 0 \ a < x \ x \ b. jumpF f (at_left x)))" define right where "right=(\f. (\x | jumpF f (at_right x) \ 0 \ a \ x \ x < b. jumpF f (at_right x)))" have "left f = 0" proof - have "jumpF f (at_left x) = 0" when "ab" for x apply (rule jumpF_eventually_const[of _ c]) unfolding eventually_at_left using assms that by auto then show ?thesis unfolding left_def by auto qed moreover have "right f = 0" proof - have "jumpF f (at_right x) = 0" when "a\x" "x real" defines "h \ (\x. f x/g x)" assumes "ag x=0) \ a\x\x\b}" and g_imp_f:"\x\{a..b}. g x=0 \ f x\0" and f_cont:"continuous_on {a..b} f" and g_cont:"continuous_on {a..b} g" shows "cindexE a b h = jumpF h (at_right a) + cindex a b h - jumpF h (at_left b)" proof - define R where "R=(\S.{x. jumpF h (at_right x) \ 0 \ x\S})" define L where "L=(\S.{x. jumpF h (at_left x) \ 0 \ x\S})" define right where "right = (\S. (\x\R S. jumpF h (at_right x)))" define left where "left = (\S. (\x\L S. jumpF h (at_left x)))" have jump_gnz:"jumpF h (at_left x) = 0" "jumpF h (at_right x) = 0" "jump h x=0" when "a0" for x proof - have "isCont h x" unfolding h_def using f_cont g_cont that by (auto intro!:continuous_intros elim:continuous_on_interior) then show "jumpF h (at_left x) = 0" "jumpF h (at_right x) = 0" "jump h x=0" using jumpF_not_infinity jump_not_infinity unfolding continuous_at_split by auto qed have finite_jFs:"finite_jumpFs h a b" proof - define S where "S=(\s. {x. (jumpF h (at_left x) \ 0 \ jumpF h (at_right x) \ 0) \ x\s})" note jump_gnz then have "S {a<.. {x. (f x=0\g x=0) \ a\x\x\b}" unfolding S_def by auto then have "finite (S {a<.. S {a,b}" unfolding S_def using \a by auto ultimately have "finite (S {a..b})" by auto then show ?thesis unfolding S_def finite_jumpFs_def by auto qed have "cindexE a b h = right {a.. R {a..a unfolding R_def by auto moreover have "R {a..x. jumpF h (at_right x)"] unfolding right_def by simp qed moreover have "left {a<..b} = jumpF h (at_left b) + left {a<.. L {a<..b}" using False \a unfolding L_def by auto moreover have "L {a<..b} - {b} = L {a<..x. jumpF h (at_left x)"] unfolding left_def by simp qed ultimately show ?thesis by simp qed also have "... = jumpF h (at_right a) + cindex a b h - jumpF h (at_left b)" proof - define S where "S={x. g x=0 \ a < x \ x < b}" have "right {a<..x. jumpF h (at_right x)) S" unfolding right_def S_def R_def apply (rule sum.mono_neutral_left) subgoal using finite_fg by (auto elim:rev_finite_subset) subgoal using jump_gnz by auto subgoal by auto done moreover have "left {a<..x. jumpF h (at_left x)) S" unfolding left_def S_def L_def apply (rule sum.mono_neutral_left) subgoal using finite_fg by (auto elim:rev_finite_subset) subgoal using jump_gnz by auto subgoal by auto done ultimately have "right {a<..x. jumpF h (at_right x) - jumpF h (at_left x)) S" by (simp add: sum_subtractf) also have "... = sum (\x. of_int(jump h x)) S" proof (rule sum.cong) fix x assume "x\S" define hr where "hr = sgnx h (at_right x)" define hl where "hl = sgnx h (at_left x)" have "h sgnx_able (at_left x)" "hr\0" "h sgnx_able (at_right x)" "hl\0" proof - have "finite {t. h t = 0 \ a < t \ t < b}" using finite_fg unfolding h_def by (auto elim!:rev_finite_subset) moreover have "continuous_on ({a<.. a < x \ x < b}) h" unfolding h_def using f_cont g_cont by (auto intro!: continuous_intros elim:continuous_on_subset) moreover have "finite {x. g x = 0 \ a < x \ x < b}" using finite_fg by (auto elim!:rev_finite_subset) moreover have " x \ {a<..x\S\ unfolding S_def by auto ultimately show "h sgnx_able (at_left x)" "hl\0" "h sgnx_able (at_right x)" "hr\0" using finite_sgnx_at_left_at_right[of h a b "{x. g x=0 \ ax h) x" proof - have "f x\0" using \x\S\ g_imp_f unfolding S_def by auto then show ?thesis using f_cont g_cont \x\S\ unfolding h_def S_def by (auto simp add:comp_def intro!:continuous_intros elim:continuous_on_interior) qed ultimately show "jumpF h (at_right x) - jumpF h (at_left x) = real_of_int (jump h x)" using jump_jumpF[of x h] \hr\0\ \hl\0\ by auto qed auto also have "... = cindex a b h" unfolding cindex_def of_int_sum S_def apply (rule sum.mono_neutral_cong_right) using jump_gnz finite_fg by (auto elim:rev_finite_subset) finally show ?thesis by simp qed finally show ?thesis . qed subsection \Cauchy index along a path\ \\Deprecated, use "cindex\_pathE" if possible\ definition cindex_path::"(real \ complex) \ complex \ int" where "cindex_path g z = cindex 0 1 (\t. Im (g t - z) / Re (g t - z))" definition cindex_pathE::"(real \ complex) \ complex \ real" where "cindex_pathE g z = cindexE 0 1 (\t. Im (g t - z) / Re (g t - z))" lemma cindex_pathE_point: "cindex_pathE (linepath a a) b = 0" unfolding cindex_pathE_def by (simp add:cindexE_constI) lemma cindex_path_reversepath: "cindex_path (reversepath g) z = - cindex_path g z" proof - define f where "f=(\t. Im (g t - z) / Re (g t - z))" define f' where "f'=(\t. Im (reversepath g t - z) / Re (reversepath g t - z))" have "f o (\t. 1 - t) = f'" unfolding f_def f'_def comp_def reversepath_def by auto then have "cindex 0 1 f' = - cindex 0 1 f" using cindex_linear_comp[of "-1" 0 1 f 1,simplified] by simp then show ?thesis unfolding cindex_path_def apply (fold f_def f'_def) by simp qed lemma cindex_pathE_reversepath: "cindex_pathE (reversepath g) z = -cindex_pathE g z" using cindexE_linear_comp[of "-1" 0 1 "\t. (Im (g t) - Im z) / (Re (g t) - Re z)" 1] by (simp add: cindex_pathE_def reversepath_def o_def) lemma cindex_pathE_reversepath': "cindex_pathE g z = -cindex_pathE (reversepath g) z" using cindexE_linear_comp[of "-1" 0 1 "\t. (Im (g t) - Im z) / (Re (g t) - Re z)" 1] by (simp add: cindex_pathE_def reversepath_def o_def) lemma cindex_pathE_joinpaths: assumes g1:"finite_ReZ_segments g1 z" and g2: "finite_ReZ_segments g2 z" and "path g1" "path g2" "pathfinish g1 = pathstart g2" shows "cindex_pathE (g1+++g2) z = cindex_pathE g1 z + cindex_pathE g2 z" proof - define f where "f = (\g (t::real). Im (g t - z) / Re (g t - z))" have "cindex_pathE (g1 +++ g2) z = cindexE 0 1 (f (g1+++g2))" unfolding cindex_pathE_def f_def by auto also have "... = cindexE 0 (1/2) (f (g1+++g2)) + cindexE (1/2) 1 (f (g1+++g2))" proof (rule cindexE_combine) show "finite_jumpFs (f (g1 +++ g2)) 0 1" unfolding f_def apply (rule finite_ReZ_segments_imp_jumpFs) subgoal using finite_ReZ_segments_joinpaths[OF g1 g2] assms(3-5) . subgoal using path_join_imp[OF \path g1\ \path g2\ \pathfinish g1=pathstart g2\] . done qed auto also have "... = cindex_pathE g1 z + cindex_pathE g2 z" proof - have "cindexE 0 (1/2) (f (g1+++g2)) = cindex_pathE g1 z" proof - have "cindexE 0 (1/2) (f (g1+++g2)) = cindexE 0 (1/2) (f g1 o ((*) 2))" apply (rule cindexE_cong) unfolding comp_def joinpaths_def f_def by auto also have "... = cindexE 0 1 (f g1)" using cindexE_linear_comp[of 2 0 "1/2" _ 0,simplified] by simp also have "... = cindex_pathE g1 z" unfolding cindex_pathE_def f_def by auto finally show ?thesis . qed moreover have "cindexE (1/2) 1 (f (g1+++g2)) = cindex_pathE g2 z" proof - have "cindexE (1/2) 1 (f (g1+++g2)) = cindexE (1/2) 1 (f g2 o (\x. 2*x - 1))" apply (rule cindexE_cong) unfolding comp_def joinpaths_def f_def by auto also have "... = cindexE 0 1 (f g2)" using cindexE_linear_comp[of 2 "1/2" 1 _ "-1",simplified] by simp also have "... = cindex_pathE g2 z" unfolding cindex_pathE_def f_def by auto finally show ?thesis . qed ultimately show ?thesis by simp qed finally show ?thesis . qed lemma cindex_pathE_constI: assumes "\t. \0 \ g t=c" shows "cindex_pathE g z = 0" unfolding cindex_pathE_def apply (rule cindexE_constI) using assms by auto lemma cindex_pathE_subpath_combine: assumes g:"finite_ReZ_segments g z"and "path g" and "0\a" "a\b" "b\c" "c\1" shows "cindex_pathE (subpath a b g) z + cindex_pathE (subpath b c g) z = cindex_pathE (subpath a c g) z" proof - define f where "f = (\t. Im (g t - z) / Re (g t - z))" have ?thesis when "a=b" proof - have "cindex_pathE (subpath a b g) z = 0" apply (rule cindex_pathE_constI) using that unfolding subpath_def by auto then show ?thesis using that by auto qed moreover have ?thesis when "b=c" proof - have "cindex_pathE (subpath b c g) z = 0" apply (rule cindex_pathE_constI) using that unfolding subpath_def by auto then show ?thesis using that by auto qed moreover have ?thesis when "a\b" "b\c" proof - have [simp]:"aa\b\ \b\c\ by auto have "cindex_pathE (subpath a b g) z = cindexE a b f" proof - have "cindex_pathE (subpath a b g) z = cindexE 0 1 (f \ (\x. (b - a) * x + a))" unfolding cindex_pathE_def f_def comp_def subpath_def by auto also have "... = cindexE a b f" using cindexE_linear_comp[of "b-a" 0 1 f a,simplified] that(1) by auto finally show ?thesis . qed moreover have "cindex_pathE (subpath b c g) z = cindexE b c f" proof - have "cindex_pathE (subpath b c g) z = cindexE 0 1 (f \ (\x. (c - b) * x + b))" unfolding cindex_pathE_def f_def comp_def subpath_def by auto also have "... = cindexE b c f" using cindexE_linear_comp[of "c-b" 0 1 f b,simplified] that(2) by auto finally show ?thesis . qed moreover have "cindex_pathE (subpath a c g) z = cindexE a c f" proof - have "cindex_pathE (subpath a c g) z = cindexE 0 1 (f \ (\x. (c - a) * x + a))" unfolding cindex_pathE_def f_def comp_def subpath_def by auto also have "... = cindexE a c f" using cindexE_linear_comp[of "c-a" 0 1 f a,simplified] \a by auto finally show ?thesis . qed moreover have "cindexE a b f + cindexE b c f = cindexE a c f " proof - have "finite_jumpFs f a c" using finite_ReZ_segments_imp_jumpFs[OF g \path g\] \0\a\ \c\1\ unfolding f_def by (elim finite_jumpFs_subE,auto) then show ?thesis using cindexE_linear_comp cindexE_combine[OF _ \a\b\ \b\c\] by auto qed ultimately show ?thesis by auto qed ultimately show ?thesis by blast qed lemma cindex_pathE_shiftpath: assumes "finite_ReZ_segments g z" "s\{0..1}" "path g" and loop:"pathfinish g = pathstart g" shows "cindex_pathE (shiftpath s g) z = cindex_pathE g z" proof - define f where "f=(\g t. Im (g (t::real) - z) / Re (g t - z))" have "cindex_pathE (shiftpath s g) z = cindexE 0 1 (f (shiftpath s g))" unfolding cindex_pathE_def f_def by simp also have "... = cindexE 0 (1-s) (f (shiftpath s g)) + cindexE (1-s) 1 (f (shiftpath s g))" proof (rule cindexE_combine) have "finite_ReZ_segments (shiftpath s g) z" using finite_ReZ_segments_shiftpah[OF assms] . from finite_ReZ_segments_imp_jumpFs[OF this] path_shiftpath[OF \path g\ loop \s\{0..1}\] show "finite_jumpFs (f (shiftpath s g)) 0 1" unfolding f_def by simp show "0 \ 1 - s" "1 - s \ 1" using \s\{0..1}\ by auto qed also have "... = cindexE 0 s (f g) + cindexE s 1 (f g)" proof - have "cindexE 0 (1-s) (f (shiftpath s g)) = cindexE s 1 (f g)" proof - have "cindexE 0 (1-s) (f (shiftpath s g)) = cindexE 0 (1-s) ((f g) o (\t. t+s))" apply (rule cindexE_cong) unfolding shiftpath_def f_def using \s\{0..1}\ by (auto simp add:algebra_simps) also have "...= cindexE s 1 (f g)" using cindexE_linear_comp[of 1 0 "1-s" "f g" s,simplified] . finally show ?thesis . qed moreover have "cindexE (1 - s) 1 (f (shiftpath s g)) = cindexE 0 s (f g)" proof - have "cindexE (1 - s) 1 (f (shiftpath s g)) = cindexE (1-s) 1 ((f g) o (\t. t+s-1))" apply (rule cindexE_cong) unfolding shiftpath_def f_def using \s\{0..1}\ by (auto simp add:algebra_simps) also have "... = cindexE 0 s (f g)" using cindexE_linear_comp[of 1 "1-s" 1 "f g" "s-1",simplified] by (simp add:algebra_simps) finally show ?thesis . qed ultimately show ?thesis by auto qed also have "... = cindexE 0 1 (f g)" proof (rule cindexE_combine[symmetric]) show "finite_jumpFs (f g) 0 1" using finite_ReZ_segments_imp_jumpFs[OF assms(1,3)] unfolding f_def by simp show "0 \ s" "s\1" using \s\{0..1}\ by auto qed also have "... = cindex_pathE g z" unfolding cindex_pathE_def f_def by simp finally show ?thesis . qed subsection \Cauchy's Index Theorem\ theorem winding_number_cindex_pathE_aux: fixes g::"real \ complex" assumes "finite_ReZ_segments g z" and "valid_path g" "z \ path_image g" and Re_ends:"Re (g 1) = Re z" "Re (g 0) = Re z" shows "2 * Re(winding_number g z) = - cindex_pathE g z" using assms proof (induct rule:finite_ReZ_segments_induct) case (sub0 g z) have "winding_number (subpath 0 0 g) z = 0" using \z \ path_image (subpath 0 0 g)\ unfolding subpath_refl by (auto intro!: winding_number_trivial) moreover have "cindex_pathE (subpath 0 0 g) z = 0" unfolding subpath_def by (auto intro:cindex_pathE_constI) ultimately show ?case by auto next case (subEq s g z) have Re_winding_0:"Re(winding_number h z) = 0" when Re_const:"\t\{0..1}. Re (h t) = Re z" and "valid_path h" "z\path_image h" for h proof - have "Re (winding_number (\t. h t - z) 0) = (Im (Ln (pathfinish (\t. h t - z))) - Im (Ln (pathstart (\t. h t - z)))) / (2 * pi)" apply (rule Re_winding_number_half_right[of _ 0,simplified]) using Re_const \valid_path h\ \z \ path_image h\ apply auto by (metis (no_types, hide_lams) add.commute imageE le_add_same_cancel1 order_refl path_image_def plus_complex.simps(1)) moreover have "Im (Ln (h 1 - z)) = Im (Ln (h 0 - z))" proof - define z0 where "z0 = h 0 - z" define z1 where "z1 = h 1 - z" have [simp]: "z0\0" "z1\0" "Re z0=0" "Re z1=0" using \z \ path_image h\ that(1) unfolding z1_def z0_def path_image_def by auto have ?thesis when [simp]: "Im z0>0" "Im z1>0" apply (fold z1_def z0_def) using Im_Ln_eq_pi_half[of z1] Im_Ln_eq_pi_half[of z0] by auto moreover have ?thesis when [simp]: "Im z0<0" "Im z1<0" apply (fold z1_def z0_def) using Im_Ln_eq_pi_half[of z1] Im_Ln_eq_pi_half[of z0] by auto moreover have False when "Im z0\0" "Im z1\0" proof - define f where "f=(\t. Im (h t - z))" have "\x\0. x \ 1 \ f x = 0" apply (rule IVT2'[of f 1 0 0]) using that valid_path_imp_path[OF \valid_path h\] unfolding f_def z0_def z1_def path_def by (auto intro:continuous_intros) then show False using Re_const \z \ path_image h\ unfolding f_def by (metis atLeastAtMost_iff complex_surj image_eqI minus_complex.simps(2) path_defs(4) right_minus_eq) qed moreover have False when "Im z0\0" "Im z1\0" proof - define f where "f=(\t. Im (h t - z))" have "\x\0. x \ 1 \ f x = 0" apply (rule IVT') using that valid_path_imp_path[OF \valid_path h\] unfolding f_def z0_def z1_def path_def by (auto intro:continuous_intros) then show False using Re_const \z \ path_image h\ unfolding f_def by (metis atLeastAtMost_iff complex_surj image_eqI minus_complex.simps(2) path_defs(4) right_minus_eq) qed ultimately show ?thesis by argo qed ultimately have "Re (winding_number (\t. h t - z) 0) = 0" unfolding pathfinish_def pathstart_def by auto then show ?thesis using winding_number_offset by auto qed have ?case when "s = 0" proof - have *: "\t\{0..1}. Re (g t) = Re z" using \\t\{s<..<1}. Re (g t) = Re z\ \Re (g 1) = Re z\ \Re (g 0) = Re z\ \s=0\ by force have "Re(winding_number g z) = 0" by (rule Re_winding_0[OF * \valid_path g\ \z \ path_image g\]) moreover have "cindex_pathE g z = 0" unfolding cindex_pathE_def apply (rule cindexE_constI) using * by auto ultimately show ?thesis by auto qed moreover have ?case when "s\0" proof - define g1 where "g1 = subpath 0 s g" define g2 where "g2 = subpath s 1 g" have "path g" "s>0" using valid_path_imp_path[OF \valid_path g\] that \s\{0..<1}\ by auto have "2 * Re (winding_number g z) = 2*Re (winding_number g1 z) + 2*Re (winding_number g2 z)" apply (subst winding_number_subpath_combine[OF \path g\ \z\path_image g\,of 0 s 1 ,simplified,symmetric]) using \s\{0..<1}\ unfolding g1_def g2_def by auto also have "... = - cindex_pathE g1 z - cindex_pathE g2 z" proof - have "2*Re (winding_number g1 z) = - cindex_pathE g1 z" unfolding g1_def apply (rule subEq.hyps(5)) subgoal using subEq.hyps(1) subEq.prems(1) valid_path_subpath by fastforce subgoal by (meson Path_Connected.path_image_subpath_subset atLeastAtMost_iff atLeastLessThan_iff less_eq_real_def subEq(7) subEq.hyps(1) subEq.prems(1) subsetCE valid_path_imp_path zero_le_one) subgoal by (metis Groups.add_ac(2) add_0_left diff_zero mult.right_neutral subEq(2) subEq(9) subpath_def) subgoal by (simp add: subEq.prems(4) subpath_def) done moreover have "2*Re (winding_number g2 z) = - cindex_pathE g2 z" proof - have *: "\t\{0..1}. Re (g2 t) = Re z" proof fix t::real assume "t\{0..1}" have "Re (g2 t) = Re z" when "t=0 \ t=1" using that unfolding g2_def by (metis \s \ 0\ add.left_neutral diff_add_cancel mult.commute mult.left_neutral mult_zero_left subEq.hyps(2) subEq.prems(3) subpath_def) moreover have "Re (g2 t) = Re z" when "t\{0<..<1}" proof - define t' where "t'=(1 - s) * t + s" then have "t'\{s<..<1}" using that \s\{0..<1}\ unfolding t'_def apply auto by (sos "((((A<0 * (A<1 * A<2)) * R<1) + ((A<=1 * (A<0 * R<1)) * (R<1 * [1]^2))))") then have "Re (g t') = Re z" using \\t\{s<..<1}. Re (g t) = Re z\ by auto then show ?thesis unfolding g2_def subpath_def t'_def . qed ultimately show "Re (g2 t) = Re z" using \t\{0..1}\ by fastforce qed have "Re(winding_number g2 z) = 0" apply (rule Re_winding_0[OF *]) subgoal using g2_def subEq.hyps(1) subEq.prems(1) valid_path_subpath by fastforce subgoal by (metis (no_types, hide_lams) Path_Connected.path_image_subpath_subset atLeastAtMost_iff atLeastLessThan_iff g2_def less_eq_real_def subEq.hyps(1) subEq.prems(1) subEq.prems(2) subsetCE valid_path_imp_path zero_le_one) done moreover have "cindex_pathE g2 z = 0" unfolding cindex_pathE_def apply (rule cindexE_constI) using * by auto ultimately show ?thesis by auto qed ultimately show ?thesis by auto qed also have "... = - cindex_pathE g z" proof - have "finite_ReZ_segments g z" unfolding finite_ReZ_segments_def apply (rule finite_Psegments.insertI_1[of s]) subgoal using \s \ {0..<1}\ by auto subgoal using \s = 0 \ Re (g s) = Re z\ by auto subgoal using \\t\{s<..<1}. Re (g t) = Re z\ by auto subgoal proof - have "finite_Psegments (\t. Re (g (s * t)) = Re z) 0 1" using \finite_ReZ_segments (subpath 0 s g) z\ unfolding subpath_def finite_ReZ_segments_def by auto from finite_Psegments_pos_linear[of _ "1/s" 0 0 s,simplified,OF this] show "finite_Psegments (\t. Re (g t - z) = 0) 0 s" using \s>0\ unfolding comp_def by auto qed done then show ?thesis using cindex_pathE_subpath_combine[OF _ \path g\,of z 0 s 1,folded g1_def g2_def,simplified] \s\{0..<1}\ by auto qed finally show ?thesis . qed ultimately show ?case by auto next case (subNEq s g z) have Re_winding:"2*Re(winding_number h z) = jumpF_pathfinish h z - jumpF_pathstart h z" when Re_neq:"\t\{0<..<1}. Re (h t) \ Re z" and "Re (h 0) = Re z" "Re (h 1) = Re z" and "valid_path h" "z\path_image h" for h proof - have Re_winding_pos: "2*Re(winding_number h0 0) = jumpF_pathfinish h0 0 - jumpF_pathstart h0 0" when Re_gt:"\t\{0<..<1}. Re (h0 t) > 0" and "Re (h0 0) = 0" "Re (h0 1) = 0" and "valid_path h0" "0\path_image h0" for h0 proof - define f where "f \ (\(t::real). Im(h0 t) / Re (h0 t))" define ln0 where "ln0 = Im (Ln (h0 0)) / pi" define ln1 where "ln1 = Im (Ln (h0 1)) / pi" have "path h0" using \valid_path h0\ valid_path_imp_path by auto have "h0 0\0" "h0 1\0" using path_defs(4) that(5) by fastforce+ have "ln1 = jumpF_pathfinish h0 0" proof - have sgnx_at_left:"((\x. Re (h0 x)) has_sgnx 1) (at_left 1)" unfolding has_sgnx_def eventually_at_left using \\p\{0<..<1}. Re (h0 p) > 0\ by (intro exI[where x=0],auto) have cont:"continuous (at_left 1) (\t. Im (h0 t))" "continuous (at_left 1) (\t. Re (h0 t))" using \path h0\ unfolding path_def by (auto intro:continuous_on_at_left[of 0 1] continuous_intros) have ?thesis when "Im (h0 1) > 0" proof - have "ln1 = 1/2" using Im_Ln_eq_pi_half[OF \h0 1\0\] that \Re (h0 1) = 0\ unfolding ln1_def by auto moreover have "jumpF_pathfinish h0 0 = 1/2" proof - have "filterlim f at_top (at_left 1)" unfolding f_def apply (subst filterlim_divide_at_bot_at_top_iff[of _ " Im (h0 1)"]) using \Re(h0 1) = 0\ sgnx_at_left cont that unfolding continuous_within by auto then show ?thesis unfolding jumpF_pathfinish_def jumpF_def f_def by auto qed ultimately show ?thesis by auto qed moreover have ?thesis when "Im (h0 1) < 0" proof - have "ln1 = - 1/2" using Im_Ln_eq_pi_half[OF \h0 1\0\] that \Re (h0 1) = 0\ unfolding ln1_def by auto moreover have "jumpF_pathfinish h0 0 = - 1/2" proof - have "((\x. Re (h0 x)) has_sgnx - sgn (Im (h0 1))) (at_left 1)" using sgnx_at_left that by auto then have "filterlim f at_bot (at_left 1)" unfolding f_def using cont that apply (subst filterlim_divide_at_bot_at_top_iff[of _ " Im (h0 1)"]) unfolding continuous_within using \Re(h0 1) = 0\ by auto then show ?thesis unfolding jumpF_pathfinish_def jumpF_def f_def by auto qed ultimately show ?thesis by auto qed moreover have "Im (h0 1)\0" using \h0 1\0\ \Re (h0 1) = 0\ using complex.expand by auto ultimately show ?thesis by linarith qed moreover have "ln0 = jumpF_pathstart h0 0" proof - have sgnx_at_right:"((\x. Re (h0 x)) has_sgnx 1) (at_right 0)" unfolding has_sgnx_def eventually_at_right using \\p\{0<..<1}. Re (h0 p) > 0\ by (intro exI[where x=1],auto) have cont:"continuous (at_right 0) (\t. Im (h0 t))" "continuous (at_right 0) (\t. Re (h0 t))" using \path h0\ unfolding path_def by (auto intro:continuous_on_at_right[of 0 1] continuous_intros) have ?thesis when "Im (h0 0) > 0" proof - have "ln0 = 1/2" using Im_Ln_eq_pi_half[OF \h0 0\0\] that \Re (h0 0) = 0\ unfolding ln0_def by auto moreover have "jumpF_pathstart h0 0 = 1/2" proof - have "filterlim f at_top (at_right 0)" unfolding f_def apply (subst filterlim_divide_at_bot_at_top_iff[of _ " Im (h0 0)"]) using \Re(h0 0) = 0\ sgnx_at_right cont that unfolding continuous_within by auto then show ?thesis unfolding jumpF_pathstart_def jumpF_def f_def by auto qed ultimately show ?thesis by auto qed moreover have ?thesis when "Im (h0 0) < 0" proof - have "ln0 = - 1/2" using Im_Ln_eq_pi_half[OF \h0 0\0\] that \Re (h0 0) = 0\ unfolding ln0_def by auto moreover have "jumpF_pathstart h0 0 = - 1/2" proof - have "filterlim f at_bot (at_right 0)" unfolding f_def apply (subst filterlim_divide_at_bot_at_top_iff[of _ " Im (h0 0)"]) using \Re(h0 0) = 0\ sgnx_at_right cont that unfolding continuous_within by auto then show ?thesis unfolding jumpF_pathstart_def jumpF_def f_def by auto qed ultimately show ?thesis by auto qed moreover have "Im (h0 0)\0" using \h0 0\0\ \Re (h0 0) = 0\ using complex.expand by auto ultimately show ?thesis by linarith qed moreover have "2*Re(winding_number h0 0) = ln1 - ln0" proof - have "\p\path_image h0. 0 \ Re p" proof fix p assume "p \ path_image h0" then obtain t where t:"t\{0..1}" "p = h0 t" unfolding path_image_def by auto have "0 \ Re p" when "t=0 \ t=1" using that t \Re (h0 0) = 0\ \Re (h0 1) = 0\ by auto moreover have "0 \ Re p" when "t\{0<..<1}" using that t Re_gt[rule_format, of t] by fastforce ultimately show "0 \ Re p" using t(1) by fastforce qed from Re_winding_number_half_right[of _ 0,simplified,OF this \valid_path h0\ \0 \ path_image h0\] show ?thesis unfolding ln1_def ln0_def pathfinish_def pathstart_def by (auto simp add:field_simps) qed ultimately show ?thesis by auto qed have ?thesis when "\t\{0<..<1}. Re (h t) < Re z" proof - let ?hu= "\t. z - h t" have "2*Re(winding_number ?hu 0) = jumpF_pathfinish ?hu 0 - jumpF_pathstart ?hu 0" apply(rule Re_winding_pos) subgoal using that by auto subgoal using \Re (h 0) = Re z\ by auto subgoal using \Re (h 1) = Re z\ by auto subgoal using \valid_path h\ valid_path_offset valid_path_uminus_comp unfolding comp_def by fastforce subgoal using \z\path_image h\ by (simp add: image_iff path_defs(4)) done moreover have "winding_number ?hu 0 = winding_number h z" using winding_number_offset[of h z] winding_number_uminus_comp[of "\t. h t- z" 0,unfolded comp_def,simplified] \valid_path h\ \z\path_image h\ by auto moreover have "jumpF_pathfinish ?hu 0 = jumpF_pathfinish h z" unfolding jumpF_pathfinish_def apply (auto intro!:jumpF_cong eventuallyI) by (auto simp add:divide_simps algebra_simps) moreover have "jumpF_pathstart ?hu 0 = jumpF_pathstart h z" unfolding jumpF_pathstart_def apply (auto intro!:jumpF_cong eventuallyI) by (auto simp add:divide_simps algebra_simps) ultimately show ?thesis by auto qed moreover have ?thesis when "\t\{0<..<1}. Re (h t) > Re z" proof - let ?hu= "\t. h t - z" have "2*Re(winding_number ?hu 0) = jumpF_pathfinish ?hu 0 - jumpF_pathstart ?hu 0" apply(rule Re_winding_pos) subgoal using that by auto subgoal using \Re (h 0) = Re z\ by auto subgoal using \Re (h 1) = Re z\ by auto subgoal using \valid_path h\ valid_path_offset valid_path_uminus_comp unfolding comp_def by fastforce subgoal using \z\path_image h\ by simp done moreover have "winding_number ?hu 0 = winding_number h z" using winding_number_offset[of h z] \valid_path h\ \z\path_image h\ by auto moreover have "jumpF_pathfinish ?hu 0 = jumpF_pathfinish h z" unfolding jumpF_pathfinish_def by auto moreover have "jumpF_pathstart ?hu 0 = jumpF_pathstart h z" unfolding jumpF_pathstart_def by auto ultimately show ?thesis by auto qed moreover have "(\t\{0<..<1}. Re (h t) > Re z) \ (\t\{0<..<1}. Re (h t) < Re z)" proof (rule ccontr) assume " \ ((\t\{0<..<1}. Re z < Re (h t)) \ (\t\{0<..<1}. Re (h t) < Re z))" then obtain t1 t2 where t:"t1\{0<..<1}" "t2\{0<..<1}" "Re (h t1)\Re z" "Re (h t2)\Re z" unfolding path_image_def by auto have False when "t1\t2" proof - have "continuous_on {t1..t2} (\t. Re (h t))" using valid_path_imp_path[OF \valid_path h\] t unfolding path_def by (metis (full_types) atLeastatMost_subset_iff continuous_on_Re continuous_on_subset eucl_less_le_not_le greaterThanLessThan_iff) then obtain t' where t':"t'\t1" "t'\t2" "Re (h t') = Re z" using IVT'[of "\t. Re (h t)" t1 _ t2] t \t1\t2\ by auto then have "t'\{0<..<1}" using t by auto then have "Re (h t') \ Re z" using Re_neq by auto then show False using \Re (h t') = Re z\ by simp qed moreover have False when "t1\t2" proof - have "continuous_on {t2..t1} (\t. Re (h t))" using valid_path_imp_path[OF \valid_path h\] t unfolding path_def by (metis (full_types) atLeastatMost_subset_iff continuous_on_Re continuous_on_subset eucl_less_le_not_le greaterThanLessThan_iff) then obtain t' where t':"t'\t1" "t'\t2" "Re (h t') = Re z" using IVT2'[of "\t. Re (h t)" t1 _ t2] t \t1\t2\ by auto then have "t'\{0<..<1}" using t by auto then have "Re (h t') \ Re z" using Re_neq by auto then show False using \Re (h t') = Re z\ by simp qed ultimately show False by linarith qed ultimately show ?thesis by blast qed have index_ends:"cindex_pathE h z = jumpF_pathstart h z - jumpF_pathfinish h z" when Re_neq:"\t\{0<..<1}. Re (h t) \ Re z" and "valid_path h" for h proof - define f where "f = (\t. Im (h t - z) / Re (h t - z))" define Ri where "Ri = {x. jumpF f (at_right x) \ 0 \ 0 \ x \ x < 1}" define Le where "Le = {x. jumpF f (at_left x) \ 0 \ 0 < x \ x \ 1}" have "path h" using \valid_path h\ valid_path_imp_path by auto have jumpF_eq0: "jumpF f (at_left x) = 0" "jumpF f (at_right x) = 0" when "x\{0<..<1}" for x proof - have "Re (h x) \ Re z" using \\t\{0<..<1}. Re (h t) \ Re z\ that by blast then have "isCont f x" unfolding f_def using continuous_on_interior[OF \path h\[unfolded path_def]] that by (auto intro!: continuous_intros isCont_Im isCont_Re) then show "jumpF f (at_left x) = 0" "jumpF f (at_right x) = 0" unfolding continuous_at_split by (auto intro: jumpF_not_infinity) qed have "cindex_pathE h z = cindexE 0 1 f" unfolding cindex_pathE_def f_def by simp also have "... = sum (\x. jumpF f (at_right x)) Ri - sum (\x. jumpF f (at_left x)) Le" unfolding cindexE_def Ri_def Le_def by auto also have "... = jumpF f (at_right 0) - jumpF f (at_left 1)" proof - have "sum (\x. jumpF f (at_right x)) Ri = jumpF f (at_right 0)" proof (cases "jumpF f (at_right 0) = 0") case True hence False if "x \ Ri" for x using that by (cases "x = 0") (auto simp: jumpF_eq0 Ri_def) hence "Ri = {}" by blast then show ?thesis using True by auto next case False hence "x \ Ri \ x = 0" for x using that by (cases "x = 0") (auto simp: jumpF_eq0 Ri_def) hence "Ri = {0}" by blast then show ?thesis by auto qed moreover have "sum (\x. jumpF f (at_left x)) Le = jumpF f (at_left 1)" proof (cases "jumpF f (at_left 1) = 0") case True then have "Le = {}" unfolding Le_def using jumpF_eq0(1) greaterThanLessThan_iff by fastforce then show ?thesis using True by auto next case False then have "Le = {1}" unfolding Le_def using jumpF_eq0(1) greaterThanLessThan_iff by fastforce then show ?thesis by auto qed ultimately show ?thesis by auto qed also have "... = jumpF_pathstart h z - jumpF_pathfinish h z" unfolding jumpF_pathstart_def jumpF_pathfinish_def f_def by simp finally show ?thesis . qed have ?case when "s=0" proof - have "2 * Re (winding_number g z) = jumpF_pathfinish g z - jumpF_pathstart g z" apply (rule Re_winding) using subNEq that by auto moreover have "cindex_pathE g z = jumpF_pathstart g z - jumpF_pathfinish g z" apply (rule index_ends) using subNEq that by auto ultimately show ?thesis by auto qed moreover have ?case when "s\0" proof - define g1 where "g1 = subpath 0 s g" define g2 where "g2 = subpath s 1 g" have "path g" "s>0" using valid_path_imp_path[OF \valid_path g\] that \s\{0..<1}\ by auto have "2 * Re (winding_number g z) = 2*Re (winding_number g1 z) + 2*Re (winding_number g2 z)" apply (subst winding_number_subpath_combine[OF \path g\ \z\path_image g\,of 0 s 1 ,simplified,symmetric]) using \s\{0..<1}\ unfolding g1_def g2_def by auto also have "... = - cindex_pathE g1 z - cindex_pathE g2 z" proof - have "2*Re (winding_number g1 z) = - cindex_pathE g1 z" unfolding g1_def apply (rule subNEq.hyps(5)) subgoal using subNEq.hyps(1) subNEq.prems(1) valid_path_subpath by fastforce subgoal by (meson Path_Connected.path_image_subpath_subset atLeastAtMost_iff atLeastLessThan_iff less_eq_real_def subNEq(7) subNEq.hyps(1) subNEq.prems(1) subsetCE valid_path_imp_path zero_le_one) subgoal by (metis Groups.add_ac(2) add_0_left diff_zero mult.right_neutral subNEq(2) subNEq(9) subpath_def) subgoal by (simp add: subNEq.prems(4) subpath_def) done moreover have "2*Re (winding_number g2 z) = - cindex_pathE g2 z" proof - have *:"\t\{0<..<1}. Re (g2 t) \ Re z" proof fix t::real assume "t \ {0<..<1}" define t' where "t'=(1 - s) * t + s" have "t'\{s<..<1}" unfolding t'_def using \s\{0..<1}\ \t \ {0<..<1}\ apply (auto simp add:algebra_simps) by (sos "((((A<0 * (A<1 * A<2)) * R<1) + ((A<=1 * (A<1 * R<1)) * (R<1 * [1]^2))))") then have "Re (g t') \ Re z" using \\t\{s<..<1}. Re (g t) \ Re z\ by auto then show "Re (g2 t) \ Re z" unfolding g2_def subpath_def t'_def by auto qed have "2*Re (winding_number g2 z) = jumpF_pathfinish g2 z - jumpF_pathstart g2 z" apply (rule Re_winding[OF *]) subgoal by (metis add.commute add.right_neutral g2_def mult_zero_right subNEq.hyps(2) subpath_def that) subgoal by (simp add: \g2 \ subpath s 1 g\ subNEq.prems(3) subpath_def) subgoal using g2_def subNEq.hyps(1) subNEq.prems(1) valid_path_subpath by fastforce subgoal by (metis (no_types, hide_lams) Path_Connected.path_image_subpath_subset \path g\ atLeastAtMost_iff atLeastLessThan_iff g2_def less_eq_real_def subNEq.hyps(1) subNEq.prems(2) subsetCE zero_le_one) done moreover have "cindex_pathE g2 z = jumpF_pathstart g2 z - jumpF_pathfinish g2 z" apply (rule index_ends[OF *]) using g2_def subNEq.hyps(1) subNEq.prems(1) valid_path_subpath by fastforce ultimately show ?thesis by auto qed ultimately show ?thesis by auto qed also have "... = - cindex_pathE g z" proof - have "finite_ReZ_segments g z" unfolding finite_ReZ_segments_def apply (rule finite_Psegments.insertI_2[of s]) subgoal using \s \ {0..<1}\ by auto subgoal using \s = 0 \ Re (g s) = Re z\ by auto subgoal using \\t\{s<..<1}. Re (g t) \ Re z\ by auto subgoal proof - have "finite_Psegments (\t. Re (g (s * t)) = Re z) 0 1" using \finite_ReZ_segments (subpath 0 s g) z\ unfolding subpath_def finite_ReZ_segments_def by auto from finite_Psegments_pos_linear[of _ "1/s" 0 0 s,simplified,OF this] show "finite_Psegments (\t. Re (g t - z) = 0) 0 s" using \s>0\ unfolding comp_def by auto qed done then show ?thesis using cindex_pathE_subpath_combine[OF _ \path g\,of z 0 s 1,folded g1_def g2_def,simplified] \s\{0..<1}\ by auto qed finally show ?thesis . qed ultimately show ?case by auto qed theorem winding_number_cindex_pathE: fixes g::"real \ complex" assumes "finite_ReZ_segments g z" and "valid_path g" "z \ path_image g" and loop: "pathfinish g = pathstart g" shows "winding_number g z = - cindex_pathE g z / 2" proof (rule finite_ReZ_segment_cases[OF assms(1)]) fix s assume "s \ {0..<1}" "s = 0 \ Re (g s) = Re z" and const:"\t\{s<..<1}. Re (g t) = Re z" and finite:"finite_ReZ_segments (subpath 0 s g) z" have "Re (g 1) = Re z" apply(rule continuous_constant_on_closure[of "{s<..<1}" "\t. Re(g t)"]) subgoal using valid_path_imp_path[OF \valid_path g\,unfolded path_def] \s\{0..<1}\ by (auto intro!:continuous_intros continuous_Re elim:continuous_on_subset) subgoal using const by auto subgoal using \s\{0..<1}\ by auto done moreover then have "Re (g 0) = Re z" using loop unfolding path_defs by auto ultimately have "2 * Re (winding_number g z) = - cindex_pathE g z" using winding_number_cindex_pathE_aux[of g z] assms(1-3) by auto moreover have "winding_number g z \ \" using integer_winding_number[OF _ loop \z\path_image g\] valid_path_imp_path[OF \valid_path g\] by auto ultimately show "winding_number g z = - cindex_pathE g z / 2" by (metis add.right_neutral complex_eq complex_is_Int_iff mult_zero_right nonzero_mult_div_cancel_left of_real_0 zero_neq_numeral) next fix s assume "s \ {0..<1}" "s = 0 \ Re (g s) = Re z" and Re_neq:"\t\{s<..<1}. Re (g t) \ Re z" and finite:"finite_ReZ_segments (subpath 0 s g) z" have "path g" using \valid_path g\ valid_path_imp_path by auto let ?goal = "2 * Re (winding_number g z) = - cindex_pathE g z" have ?goal when "s=0" proof - have index_ends:"cindex_pathE h z = jumpF_pathstart h z - jumpF_pathfinish h z" when Re_neq:"\t\{0<..<1}. Re (h t) \ Re z" and "valid_path h" for h proof - define f where "f = (\t. Im (h t - z) / Re (h t - z))" define Ri where "Ri = {x. jumpF f (at_right x) \ 0 \ 0 \ x \ x < 1}" define Le where "Le = {x. jumpF f (at_left x) \ 0 \ 0 < x \ x \ 1}" have "path h" using \valid_path h\ valid_path_imp_path by auto have jumpF_eq0: "jumpF f (at_left x) = 0" "jumpF f (at_right x) = 0" when "x\{0<..<1}" for x proof - have "Re (h x) \ Re z" using \\t\{0<..<1}. Re (h t) \ Re z\ that by blast then have "isCont f x" unfolding f_def using continuous_on_interior[OF \path h\[unfolded path_def]] that by (auto intro!: continuous_intros isCont_Im isCont_Re) then show "jumpF f (at_left x) = 0" "jumpF f (at_right x) = 0" unfolding continuous_at_split by (auto intro: jumpF_not_infinity) qed have "cindex_pathE h z = cindexE 0 1 f" unfolding cindex_pathE_def f_def by simp also have "... = sum (\x. jumpF f (at_right x)) Ri - sum (\x. jumpF f (at_left x)) Le" unfolding cindexE_def Ri_def Le_def by auto also have "... = jumpF f (at_right 0) - jumpF f (at_left 1)" proof - have "sum (\x. jumpF f (at_right x)) Ri = jumpF f (at_right 0)" proof (cases "jumpF f (at_right 0) = 0") case True hence False if "x \ Ri" for x using that by (cases "x = 0") (auto simp: jumpF_eq0 Ri_def) hence "Ri = {}" by blast then show ?thesis using True by auto next case False hence "x \ Ri \ x = 0" for x using that by (cases "x = 0") (auto simp: jumpF_eq0 Ri_def) then have "Ri = {0}" by blast then show ?thesis by auto qed moreover have "sum (\x. jumpF f (at_left x)) Le = jumpF f (at_left 1)" proof (cases "jumpF f (at_left 1) = 0") case True then have "Le = {}" unfolding Le_def using jumpF_eq0(1) greaterThanLessThan_iff by fastforce then show ?thesis using True by auto next case False then have "Le = {1}" unfolding Le_def using jumpF_eq0(1) greaterThanLessThan_iff by fastforce then show ?thesis by auto qed ultimately show ?thesis by auto qed also have "... = jumpF_pathstart h z - jumpF_pathfinish h z" unfolding jumpF_pathstart_def jumpF_pathfinish_def f_def by simp finally show ?thesis . qed define fI where "fI=(\t. Im (g t - z))" define fR where "fR=(\t. Re (g t - z))" have fI: "(fI \ fI 0) (at_right 0)" "(fI \ fI 1) (at_left 1)" proof - have "continuous (at_right 0) fI" apply (rule continuous_on_at_right[of _ 1]) using \path g\ unfolding fI_def path_def by (auto intro:continuous_intros) then show "(fI \ fI 0) (at_right 0)" by (simp add: continuous_within) next have "continuous (at_left 1) fI" apply (rule continuous_on_at_left[of 0]) using \path g\ unfolding fI_def path_def by (auto intro:continuous_intros) then show "(fI \ fI 1) (at_left 1)" by (simp add: continuous_within) qed have fR: "(fR \ 0) (at_right 0)" "(fR \ 0) (at_left 1)" when "Re (g 0) = Re z" proof - have "continuous (at_right 0) fR" apply (rule continuous_on_at_right[of _ 1]) using \path g\ unfolding fR_def path_def by (auto intro:continuous_intros) then show "(fR \ 0) (at_right 0)" using that unfolding fR_def by (simp add: continuous_within) next have "continuous (at_left 1) fR" apply (rule continuous_on_at_left[of 0]) using \path g\ unfolding fR_def path_def by (auto intro:continuous_intros) then show "(fR \ 0) (at_left 1)" using that loop unfolding fR_def path_defs by (simp add: continuous_within) qed have "(\t\{0<..<1}. Re (g t) > Re z) \ (\t\{0<..<1}. Re (g t) < Re z)" proof (rule ccontr) assume " \ ((\t\{0<..<1}. Re z < Re (g t)) \ (\t\{0<..<1}. Re (g t) < Re z))" then obtain t1 t2 where t:"t1\{0<..<1}" "t2\{0<..<1}" "Re (g t1)\Re z" "Re (g t2)\Re z" unfolding path_image_def by auto have False when "t1\t2" proof - have "continuous_on {t1..t2} (\t. Re (g t))" using valid_path_imp_path[OF \valid_path g\] t unfolding path_def by (metis (full_types) atLeastatMost_subset_iff continuous_on_Re continuous_on_subset eucl_less_le_not_le greaterThanLessThan_iff) then obtain t' where t':"t'\t1" "t'\t2" "Re (g t') = Re z" using IVT'[of "\t. Re (g t)" t1 _ t2] t \t1\t2\ by auto then have "t'\{0<..<1}" using t by auto then have "Re (g t') \ Re z" using Re_neq \s=0\ by auto then show False using \Re (g t') = Re z\ by simp qed moreover have False when "t1\t2" proof - have "continuous_on {t2..t1} (\t. Re (g t))" using valid_path_imp_path[OF \valid_path g\] t unfolding path_def by (metis (full_types) atLeastatMost_subset_iff continuous_on_Re continuous_on_subset eucl_less_le_not_le greaterThanLessThan_iff) then obtain t' where t':"t'\t1" "t'\t2" "Re (g t') = Re z" using IVT2'[of "\t. Re (g t)" t1 _ t2] t \t1\t2\ by auto then have "t'\{0<..<1}" using t by auto then have "Re (g t') \ Re z" using Re_neq \s=0\ by auto then show False using \Re (g t') = Re z\ by simp qed ultimately show False by linarith qed moreover have ?thesis when Re_pos:"\t\{0<..<1}. Re (g t) > Re z" proof - have "Re (winding_number g z) = 0" proof - have "\p\path_image g. Re z \ Re p" proof fix p assume "p \ path_image g" then obtain t where "0\t" "t\1" "p = g t" unfolding path_image_def by auto have "Re z \ Re (g t)" apply (rule continuous_ge_on_closure[of "{0<..<1}" "\t. Re (g t)" t "Re z",simplified]) subgoal using valid_path_imp_path[OF \valid_path g\,unfolded path_def] by (auto intro:continuous_intros) subgoal using \0\t\ \t\1\ by auto subgoal for x using that[rule_format,of x] by auto done then show "Re z \ Re p" using \p = g t\ by auto qed from Re_winding_number_half_right[OF this \valid_path g\ \z\path_image g\] loop show ?thesis by auto qed moreover have "cindex_pathE g z = 0" proof - have "cindex_pathE g z = jumpF_pathstart g z - jumpF_pathfinish g z" using index_ends[OF _ \valid_path g\] Re_neq \s=0\ by auto moreover have "jumpF_pathstart g z = jumpF_pathfinish g z" when "Re (g 0) \ Re z" proof - have "jumpF_pathstart g z = 0" using jumpF_pathstart_eq_0[OF \path g\] that unfolding path_defs by auto moreover have "jumpF_pathfinish g z=0" using jumpF_pathfinish_eq_0[OF \path g\] that loop unfolding path_defs by auto ultimately show ?thesis by auto qed moreover have "jumpF_pathstart g z = jumpF_pathfinish g z" when "Re (g 0) = Re z" proof - have [simp]:"(fR has_sgnx 1) (at_right 0)" unfolding fR_def has_sgnx_def eventually_at_right apply (rule exI[where x=1]) using Re_pos by auto have [simp]:"(fR has_sgnx 1) (at_left 1)" unfolding fR_def has_sgnx_def eventually_at_left apply (rule exI[where x=0]) using Re_pos by auto have "fI 0\0" proof (rule ccontr) assume "\ fI 0 \ 0" then have "g 0 =z" using \Re (g 0) = Re z\ unfolding fI_def by (simp add: complex.expand) then show False using \z \ path_image g\ unfolding path_image_def by auto qed moreover have ?thesis when "fI 0>0" proof - have "jumpF_pathstart g z = 1/2" proof - have "(LIM x at_right 0. fI x / fR x :> at_top)" apply (subst filterlim_divide_at_bot_at_top_iff[of _ "fI 0"]) using that fI fR[OF \Re (g 0) = Re z\] by simp_all then show ?thesis unfolding jumpF_pathstart_def fI_def fR_def jumpF_def by auto qed moreover have "jumpF_pathfinish g z = 1/2" proof - have "fI 1>0" using loop that unfolding path_defs fI_def by auto then have "(LIM x at_left 1. fI x / fR x :> at_top)" apply (subst filterlim_divide_at_bot_at_top_iff[of _ "fI 1"]) using that fI fR[OF \Re (g 0) = Re z\] by simp_all then show ?thesis unfolding jumpF_pathfinish_def fI_def fR_def jumpF_def by auto qed ultimately show ?thesis by simp qed moreover have ?thesis when "fI 0<0" proof - have "jumpF_pathstart g z = - 1/2" proof - have "(LIM x at_right 0. fI x / fR x :> at_bot)" apply (subst filterlim_divide_at_bot_at_top_iff[of _ "fI 0"]) using that fI fR[OF \Re (g 0) = Re z\] by simp_all then show ?thesis unfolding jumpF_pathstart_def fI_def fR_def jumpF_def by auto qed moreover have "jumpF_pathfinish g z = - 1/2" proof - have "fI 1<0" using loop that unfolding path_defs fI_def by auto then have "(LIM x at_left 1. fI x / fR x :> at_bot)" apply (subst filterlim_divide_at_bot_at_top_iff[of _ "fI 1"]) using that fI fR[OF \Re (g 0) = Re z\] by simp_all then show ?thesis unfolding jumpF_pathfinish_def fI_def fR_def jumpF_def by auto qed ultimately show ?thesis by simp qed ultimately show ?thesis by linarith qed ultimately show ?thesis by auto qed ultimately show ?thesis by auto qed moreover have ?thesis when Re_neg:"\t\{0<..<1}. Re (g t) < Re z" proof - have "Re (winding_number g z) = 0" proof - have "\p\path_image g. Re z \ Re p" proof fix p assume "p \ path_image g" then obtain t where "0\t" "t\1" "p = g t" unfolding path_image_def by auto have "Re z \ Re (g t)" apply (rule continuous_le_on_closure[of "{0<..<1}" "\t. Re (g t)" t "Re z",simplified]) subgoal using valid_path_imp_path[OF \valid_path g\,unfolded path_def] by (auto intro:continuous_intros) subgoal using \0\t\ \t\1\ by auto subgoal for x using that[rule_format,of x] by auto done then show "Re z \ Re p" using \p = g t\ by auto qed from Re_winding_number_half_left[OF this \valid_path g\ \z\path_image g\] loop show ?thesis by auto qed moreover have "cindex_pathE g z = 0" proof - have "cindex_pathE g z = jumpF_pathstart g z - jumpF_pathfinish g z" using index_ends[OF _ \valid_path g\] Re_neq \s=0\ by auto moreover have "jumpF_pathstart g z = jumpF_pathfinish g z" when "Re (g 0) \ Re z" proof - have "jumpF_pathstart g z = 0" using jumpF_pathstart_eq_0[OF \path g\] that unfolding path_defs by auto moreover have "jumpF_pathfinish g z=0" using jumpF_pathfinish_eq_0[OF \path g\] that loop unfolding path_defs by auto ultimately show ?thesis by auto qed moreover have "jumpF_pathstart g z = jumpF_pathfinish g z" when "Re (g 0) = Re z" proof - have [simp]:"(fR has_sgnx - 1) (at_right 0)" unfolding fR_def has_sgnx_def eventually_at_right apply (rule exI[where x=1]) using Re_neg by auto have [simp]:"(fR has_sgnx - 1) (at_left 1)" unfolding fR_def has_sgnx_def eventually_at_left apply (rule exI[where x=0]) using Re_neg by auto have "fI 0\0" proof (rule ccontr) assume "\ fI 0 \ 0" then have "g 0 =z" using \Re (g 0) = Re z\ unfolding fI_def by (simp add: complex.expand) then show False using \z \ path_image g\ unfolding path_image_def by auto qed moreover have ?thesis when "fI 0>0" proof - have "jumpF_pathstart g z = - 1/2" proof - have "(LIM x at_right 0. fI x / fR x :> at_bot)" apply (subst filterlim_divide_at_bot_at_top_iff[of _ "fI 0"]) using that fI fR[OF \Re (g 0) = Re z\] by simp_all then show ?thesis unfolding jumpF_pathstart_def fI_def fR_def jumpF_def by auto qed moreover have "jumpF_pathfinish g z = - 1/2" proof - have "fI 1>0" using loop that unfolding path_defs fI_def by auto then have "(LIM x at_left 1. fI x / fR x :> at_bot)" apply (subst filterlim_divide_at_bot_at_top_iff[of _ "fI 1"]) using that fI fR[OF \Re (g 0) = Re z\] by simp_all then show ?thesis unfolding jumpF_pathfinish_def fI_def fR_def jumpF_def by auto qed ultimately show ?thesis by simp qed moreover have ?thesis when "fI 0<0" proof - have "jumpF_pathstart g z = 1/2" proof - have "(LIM x at_right 0. fI x / fR x :> at_top)" apply (subst filterlim_divide_at_bot_at_top_iff[of _ "fI 0"]) using that fI fR[OF \Re (g 0) = Re z\] by simp_all then show ?thesis unfolding jumpF_pathstart_def fI_def fR_def jumpF_def by auto qed moreover have "jumpF_pathfinish g z = 1/2" proof - have "fI 1<0" using loop that unfolding path_defs fI_def by auto then have "(LIM x at_left 1. fI x / fR x :> at_top)" apply (subst filterlim_divide_at_bot_at_top_iff[of _ "fI 1"]) using that fI fR[OF \Re (g 0) = Re z\] by simp_all then show ?thesis unfolding jumpF_pathfinish_def fI_def fR_def jumpF_def by auto qed ultimately show ?thesis by simp qed ultimately show ?thesis by linarith qed ultimately show ?thesis by auto qed ultimately show ?thesis by auto qed ultimately show ?thesis by auto qed moreover have ?goal when "s\0" proof - have "Re (g s) = Re z" using \s = 0 \ Re (g s) = Re z\ that by auto define g' where "g' = shiftpath s g" have "2 * Re (winding_number g' z) = - cindex_pathE g' z" proof (rule winding_number_cindex_pathE_aux) show "Re (g' 1) = Re z" "Re (g' 0) = Re z" using \Re (g s) = Re z\ \s\{0..<1}\ \s\0\ unfolding g'_def shiftpath_def by simp_all show "valid_path g'" using valid_path_shiftpath[OF \valid_path g\ loop,of s,folded g'_def] \s\{0..<1}\ by auto show "z \ path_image g'" using \s \ {0..<1}\ assms(3) g'_def loop path_image_shiftpath by fastforce show "finite_ReZ_segments g' z" using finite_ReZ_segments_shiftpah[OF \finite_ReZ_segments g z\ _ \path g\ loop] \s\{0..<1}\ unfolding g'_def by auto qed moreover have "winding_number g' z = winding_number g z" unfolding g'_def apply (rule winding_number_shiftpath[OF \path g\ \z \ path_image g\ loop]) using \s\{0..<1}\ by auto moreover have "cindex_pathE g' z = cindex_pathE g z" unfolding g'_def apply (rule cindex_pathE_shiftpath[OF \finite_ReZ_segments g z\ _ \path g\ loop]) using \s\{0..<1}\ by auto ultimately show ?thesis by auto qed ultimately have ?goal by auto moreover have "winding_number g z \ \" using integer_winding_number[OF _ loop \z\path_image g\] valid_path_imp_path[OF \valid_path g\] by auto ultimately show "winding_number g z = - cindex_pathE g z / 2" by (metis add.right_neutral complex_eq complex_is_Int_iff mult_zero_right nonzero_mult_div_cancel_left of_real_0 zero_neq_numeral) qed text \REMARK: The usual statement of Cauchy's Index theorem (i.e. Analytic Theory of Polynomials (2002): Theorem 11.1.3) is about the equality between the number of polynomial roots and the Cauchy index, which is the joint application of @{thm winding_number_cindex_pathE} and @{thm argument_principle}.\ end diff --git a/thys/Winding_Number_Eval/Missing_Algebraic.thy b/thys/Winding_Number_Eval/Missing_Algebraic.thy --- a/thys/Winding_Number_Eval/Missing_Algebraic.thy +++ b/thys/Winding_Number_Eval/Missing_Algebraic.thy @@ -1,371 +1,371 @@ (* Author: Wenda Li *) section \Some useful lemmas in algebra\ theory Missing_Algebraic imports "HOL-Computational_Algebra.Polynomial_Factorial" "HOL-Computational_Algebra.Fundamental_Theorem_Algebra" - "HOL-Analysis.Analysis" + "HOL-Complex_Analysis.Complex_Analysis" Missing_Topology "Budan_Fourier.BF_Misc" begin subsection \Misc\ lemma poly_holomorphic_on[simp]: "(poly p) holomorphic_on s" apply (rule holomorphic_onI) apply (unfold field_differentiable_def) apply (rule_tac x="poly (pderiv p) x" in exI) by (simp add:has_field_derivative_at_within) lemma order_zorder: fixes p::"complex poly" and z::complex assumes "p\0" shows "order z p = nat (zorder (poly p) z)" proof - define n where "n=nat (zorder (poly p) z)" define h where "h=zor_poly (poly p) z" have "\w. poly p w \ 0" using assms poly_all_0_iff_0 by auto then obtain r where "0 < r" "cball z r \ UNIV" and h_holo: "h holomorphic_on cball z r" and poly_prod:"(\w\cball z r. poly p w = h w * (w - z) ^ n \ h w \ 0)" using zorder_exist_zero[of "poly p" UNIV z,folded h_def] poly_holomorphic_on unfolding n_def by auto then have "h holomorphic_on ball z r" and "(\w\ball z r. poly p w = h w * (w - z) ^ n)" and "h z\0" by auto then have "order z p = n" using \p\0\ proof (induct n arbitrary:p h) case 0 then have "poly p z=h z" using \r>0\ by auto then have "poly p z\0" using \h z\0\ by auto then show ?case using order_root by blast next case (Suc n) define sn where "sn=Suc n" define h' where "h'\ \w. deriv h w * (w-z)+ sn * h w" have "(poly p has_field_derivative poly (pderiv p) w) (at w)" for w using poly_DERIV[of p w] . moreover have "(poly p has_field_derivative (h' w)*(w-z)^n ) (at w)" when "w\ball z r" for w proof (subst DERIV_cong_ev[of w w "poly p" "\w. h w * (w - z) ^ Suc n" ],simp_all) show "\\<^sub>F x in nhds w. poly p x = h x * ((x - z) * (x - z) ^ n)" unfolding eventually_nhds using Suc(3) \w\ball z r\ apply (intro exI[where x="ball z r"]) by auto next have "(h has_field_derivative deriv h w) (at w)" using \h holomorphic_on ball z r\ \w\ball z r\ holomorphic_on_imp_differentiable_at by (simp add: holomorphic_derivI) then have "((\w. h w * ((w - z) ^ sn)) has_field_derivative h' w * (w - z) ^ (sn - 1)) (at w)" unfolding h'_def apply (auto intro!: derivative_eq_intros simp add:field_simps) by (auto simp add:field_simps sn_def) then show "((\w. h w * ((w - z) * (w - z) ^ n)) has_field_derivative h' w * (w - z) ^ n) (at w)" unfolding sn_def by auto qed ultimately have "\w\ball z r. poly (pderiv p) w = h' w * (w - z) ^ n" using DERIV_unique by blast moreover have "h' holomorphic_on ball z r" unfolding h'_def using \h holomorphic_on ball z r\ by (auto intro!: holomorphic_intros) moreover have "h' z\0" unfolding h'_def sn_def using \h z \ 0\ of_nat_neq_0 by auto moreover have "pderiv p \ 0" proof assume "pderiv p = 0" obtain c where "p=[:c:]" using \pderiv p = 0\ using pderiv_iszero by blast then have "c=0" using Suc(3)[rule_format,of z] \r>0\ by auto then show False using \p\0\ using \p=[:c:]\ by auto qed ultimately have "order z (pderiv p) = n" by (auto elim: Suc.hyps) moreover have "order z p \ 0" using Suc(3)[rule_format,of z] \r>0\ order_root \p\0\ by auto ultimately show ?case using order_pderiv[OF \pderiv p \ 0\] by auto qed then show ?thesis unfolding n_def . qed lemma pcompose_pCons_0:"pcompose p [:a:] = [:poly p a:]" apply (induct p) by (auto simp add:pcompose_pCons algebra_simps) lemma pcompose_coeff_0: "coeff (pcompose p q) 0 = poly p (coeff q 0)" apply (induct p) by (auto simp add:pcompose_pCons coeff_mult) lemma poly_field_differentiable_at[simp]: "poly p field_differentiable (at x within s)" apply (unfold field_differentiable_def) apply (rule_tac x="poly (pderiv p) x" in exI) by (simp add:has_field_derivative_at_within) lemma deriv_pderiv: "deriv (poly p) = poly (pderiv p)" apply (rule ext) apply (rule DERIV_imp_deriv) using poly_DERIV . lemma lead_coeff_map_poly_nz: assumes "f (lead_coeff p) \0" "f 0=0" shows "lead_coeff (map_poly f p) = f (lead_coeff p) " proof - have "lead_coeff (Poly (map f (coeffs p))) = f (lead_coeff p)" by (metis (mono_tags, lifting) antisym assms(1) assms(2) coeff_0_degree_minus_1 coeff_map_poly degree_Poly degree_eq_length_coeffs le_degree length_map map_poly_def) then show ?thesis by (simp add: map_poly_def) qed lemma filterlim_poly_at_infinity: fixes p::"'a::real_normed_field poly" assumes "degree p>0" shows "filterlim (poly p) at_infinity at_infinity" using assms proof (induct p) case 0 then show ?case by auto next case (pCons a p) have ?case when "degree p=0" proof - obtain c where c_def:"p=[:c:]" using \degree p = 0\ degree_eq_zeroE by blast then have "c\0" using \0 < degree (pCons a p)\ by auto then show ?thesis unfolding c_def apply (auto intro!:tendsto_add_filterlim_at_infinity) apply (subst mult.commute) by (auto intro!:tendsto_mult_filterlim_at_infinity filterlim_ident) qed moreover have ?case when "degree p\0" proof - have "filterlim (poly p) at_infinity at_infinity" using that by (auto intro:pCons) then show ?thesis by (auto intro!:tendsto_add_filterlim_at_infinity filterlim_at_infinity_times filterlim_ident) qed ultimately show ?case by auto qed lemma poly_divide_tendsto_aux: fixes p::"'a::real_normed_field poly" shows "((\x. poly p x/x^(degree p)) \ lead_coeff p) at_infinity" proof (induct p) case 0 then show ?case by (auto intro:tendsto_eq_intros) next case (pCons a p) have ?case when "p=0" using that by auto moreover have ?case when "p\0" proof - define g where "g=(\x. a/(x*x^degree p))" define f where "f=(\x. poly p x/x^degree p)" have "\\<^sub>Fx in at_infinity. poly (pCons a p) x / x ^ degree (pCons a p) = g x + f x" proof (rule eventually_at_infinityI[of 1]) fix x::'a assume "norm x\1" then have "x\0" by auto then show "poly (pCons a p) x / x ^ degree (pCons a p) = g x + f x" using that unfolding g_def f_def by (auto simp add:field_simps) qed moreover have "((\x. g x+f x) \ lead_coeff (pCons a p)) at_infinity" proof - have "(g \ 0) at_infinity" unfolding g_def using filterlim_poly_at_infinity[of "monom 1 (Suc (degree p))"] apply (auto intro!:tendsto_intros tendsto_divide_0 simp add: degree_monom_eq) apply (subst filterlim_cong[where g="poly (monom 1 (Suc (degree p)))"]) by (auto simp add:poly_monom) moreover have "(f \ lead_coeff (pCons a p)) at_infinity" using pCons \p\0\ unfolding f_def by auto ultimately show ?thesis by (auto intro:tendsto_eq_intros) qed ultimately show ?thesis by (auto dest:tendsto_cong) qed ultimately show ?case by auto qed lemma filterlim_power_at_infinity: assumes "n\0" shows "filterlim (\x::'a::real_normed_field. x^n) at_infinity at_infinity" using filterlim_poly_at_infinity[of "monom 1 n"] assms apply (subst filterlim_cong[where g="poly (monom 1 n)"]) by (auto simp add:poly_monom degree_monom_eq) lemma poly_divide_tendsto_0_at_infinity: fixes p::"'a::real_normed_field poly" assumes "degree p > degree q" shows "((\x. poly q x / poly p x) \ 0 ) at_infinity" proof - define pp where "pp=(\x. x^(degree p) / poly p x)" define qq where "qq=(\x. poly q x/x^(degree q))" define dd where "dd=(\x::'a. 1/x^(degree p - degree q))" have "\\<^sub>Fx in at_infinity. poly q x / poly p x = qq x * pp x * dd x" proof (rule eventually_at_infinityI[of 1]) fix x::'a assume "norm x\1" then have "x\0" by auto then show "poly q x / poly p x = qq x * pp x * dd x" unfolding qq_def pp_def dd_def using assms by (auto simp add:field_simps divide_simps power_diff) qed moreover have "((\x. qq x * pp x * dd x) \ 0) at_infinity" proof - have "(qq \ lead_coeff q) at_infinity" unfolding qq_def using poly_divide_tendsto_aux[of q] . moreover have "(pp \ 1/lead_coeff p) at_infinity" proof - have "p\0" using assms by auto then show ?thesis unfolding pp_def using poly_divide_tendsto_aux[of p] apply (drule_tac tendsto_inverse) by (auto simp add:inverse_eq_divide) qed moreover have "(dd \ 0) at_infinity" unfolding dd_def apply (rule tendsto_divide_0) by (auto intro!: filterlim_power_at_infinity simp add:assms) ultimately show ?thesis by (auto intro:tendsto_eq_intros) qed ultimately show ?thesis by (auto dest:tendsto_cong) qed lemma lead_coeff_list_def: "lead_coeff p= (if coeffs p=[] then 0 else last (coeffs p))" by (simp add: last_coeffs_eq_coeff_degree) lemma poly_linepath_comp: fixes a::"'a::{real_normed_vector,comm_semiring_0,real_algebra_1}" shows "poly p o (linepath a b) = poly (p \\<^sub>p [:a, b-a:]) o of_real" apply rule by (auto simp add:poly_pcompose linepath_def scaleR_conv_of_real algebra_simps) lemma poly_eventually_not_zero: fixes p::"real poly" assumes "p\0" shows "eventually (\x. poly p x\0) at_infinity" proof (rule eventually_at_infinityI[of "Max (norm ` {x. poly p x=0}) + 1"]) fix x::real assume asm:"Max (norm ` {x. poly p x=0}) + 1 \ norm x" have False when "poly p x=0" proof - define S where "S=norm `{x. poly p x = 0}" have "norm x\S" using that unfolding S_def by auto moreover have "finite S" using \p\0\ poly_roots_finite unfolding S_def by blast ultimately have "norm x\Max S" by simp moreover have "Max S + 1 \ norm x" using asm unfolding S_def by simp ultimately show False by argo qed then show "poly p x \ 0" by auto qed subsection \More about @{term degree}\ lemma degree_div_less: fixes x y::"'a::field poly" assumes "degree x\0" "degree y\0" shows "degree (x div y) < degree x" proof - have "x\0" "y\0" using assms by auto define q r where "q=x div y" and "r=x mod y" have *:"eucl_rel_poly x y (q, r)" unfolding q_def r_def by (simp add: eucl_rel_poly) then have "r = 0 \ degree r < degree y" using \y\0\ unfolding eucl_rel_poly_iff by auto moreover have ?thesis when "r=0" proof - have "x = q * y" using * that unfolding eucl_rel_poly_iff by auto then have "q\0" using \x\0\ \y\0\ by auto from degree_mult_eq[OF this \y\0\] \x = q * y\ have "degree x = degree q +degree y" by auto then show ?thesis unfolding q_def using assms by auto qed moreover have ?thesis when "degree rdegree x") case True then have "q=0" unfolding q_def using div_poly_less by auto then show ?thesis unfolding q_def using assms(1) by auto next case False then have "degree x>degree r" using that by auto then have "degree x = degree (x-r)" using degree_add_eq_right[of "-r" x] by auto have "x-r = q*y" using * unfolding eucl_rel_poly_iff by auto then have "q\0" using \degree r < degree x\ by auto have "degree x = degree q +degree y" using degree_mult_eq[OF \q\0\ \y\0\] \x-r = q*y\ \degree x = degree (x-r)\ by auto then show ?thesis unfolding q_def using assms by auto qed ultimately show ?thesis by auto qed lemma map_poly_degree_eq: assumes "f (lead_coeff p) \0" shows "degree (map_poly f p) = degree p" using assms unfolding map_poly_def degree_eq_length_coeffs coeffs_Poly lead_coeff_list_def by (metis (full_types) last_conv_nth_default length_map no_trailing_unfold nth_default_coeffs_eq nth_default_map_eq strip_while_idem) lemma map_poly_degree_less: assumes "f (lead_coeff p) =0" "degree p\0" shows "degree (map_poly f p) < degree p" proof - have "length (coeffs p) >1" using \degree p\0\ by (simp add: degree_eq_length_coeffs) then obtain xs x where xs_def:"coeffs p=xs@[x]" "length xs>0" by (metis One_nat_def add.commute add_diff_cancel_left' append_Nil assms(2) degree_eq_length_coeffs length_greater_0_conv list.size(3) list.size(4) not_less_zero rev_exhaust) have "f x=0" using assms(1) by (simp add: lead_coeff_list_def xs_def(1)) have "degree (map_poly f p) = length (strip_while ((=) 0) (map f (xs@[x]))) - 1" unfolding map_poly_def degree_eq_length_coeffs coeffs_Poly by (subst xs_def,auto) also have "... = length (strip_while ((=) 0) (map f xs)) - 1" using \f x=0\ by simp also have "... \ length xs -1" using length_strip_while_le by (metis diff_le_mono length_map) also have "... < length (xs@[x]) - 1" using xs_def(2) by auto also have "... = degree p" unfolding degree_eq_length_coeffs xs_def by simp finally show ?thesis . qed lemma map_poly_degree_leq[simp]: shows "degree (map_poly f p) \ degree p" unfolding map_poly_def degree_eq_length_coeffs by (metis coeffs_Poly diff_le_mono length_map length_strip_while_le) subsection \roots / zeros of a univariate function\ definition roots_within::"('a \ 'b::zero) \ 'a set \ 'a set" where "roots_within f s = {x\s. f x = 0}" abbreviation roots::"('a \ 'b::zero) \ 'a set" where "roots f \ roots_within f UNIV" subsection \The argument principle specialised to polynomials.\ lemma argument_principle_poly: assumes "p\0" and valid:"valid_path g" and loop: "pathfinish g = pathstart g" and no_proots:"path_image g \ - proots p" shows "contour_integral g (\x. deriv (poly p) x / poly p x) = 2 * of_real pi * \ * (\x\proots p. winding_number g x * of_nat (order x p))" proof - have "contour_integral g (\x. deriv (poly p) x / poly p x) = 2 * of_real pi * \ * (\x | poly p x = 0. winding_number g x * of_int (zorder (poly p) x))" apply (rule argument_principle[of UNIV "poly p" "{}" "\_. 1" g,simplified,OF _ valid loop]) using no_proots[unfolded proots_def] by (auto simp add:poly_roots_finite[OF \p\0\] ) also have "... = 2 * of_real pi * \ * (\x\proots p. winding_number g x * of_nat (order x p))" proof - have "nat (zorder (poly p) x) = order x p" when "x\proots p" for x using order_zorder[OF \p\0\] that unfolding proots_def by auto then show ?thesis unfolding proots_def apply (auto intro!:comm_monoid_add_class.sum.cong) by (metis assms(1) nat_eq_iff2 of_nat_nat order_root) qed finally show ?thesis . qed end diff --git a/thys/Winding_Number_Eval/Missing_Analysis.thy b/thys/Winding_Number_Eval/Missing_Analysis.thy --- a/thys/Winding_Number_Eval/Missing_Analysis.thy +++ b/thys/Winding_Number_Eval/Missing_Analysis.thy @@ -1,154 +1,153 @@ (* Author: Wenda Li *) section \Some useful lemmas in analysis\ -theory Missing_Analysis imports "HOL-Analysis.Analysis" +theory Missing_Analysis + imports "HOL-Complex_Analysis.Complex_Analysis" begin subsection \More about paths\ lemma pathfinish_offset[simp]: "pathfinish (\t. g t - z) = pathfinish g - z" unfolding pathfinish_def by simp lemma pathstart_offset[simp]: "pathstart (\t. g t - z) = pathstart g - z" unfolding pathstart_def by simp lemma pathimage_offset[simp]: fixes g :: "_ \ 'b::topological_group_add" shows "p \ path_image (\t. g t - z) \ p+z \ path_image g " unfolding path_image_def by (auto simp:algebra_simps) lemma path_offset[simp]: fixes g :: "_ \ 'b::topological_group_add" shows "path (\t. g t - z) \ path g" unfolding path_def proof assume "continuous_on {0..1} (\t. g t - z)" hence "continuous_on {0..1} (\t. (g t - z) + z)" apply (rule continuous_intros) by (intro continuous_intros) then show "continuous_on {0..1} g" by auto qed (auto intro:continuous_intros) lemma not_on_circlepathI: assumes "cmod (z-z0) \ \r\" shows "z \ path_image (part_circlepath z0 r st tt)" proof (rule ccontr) assume "\ z \ path_image (part_circlepath z0 r st tt)" then have "z\path_image (part_circlepath z0 r st tt)" by simp then obtain t where "t\{0..1}" and *:"z = z0 + r * exp (\ * (linepath st tt t))" unfolding path_image_def image_def part_circlepath_def by blast define \ where "\ = linepath st tt t" then have "z-z0 = r * exp (\ * \)" using * by auto then have "cmod (z-z0) = cmod (r * exp (\ * \))" by auto also have "\ = \r\ * cmod (exp (\ * \))" by (simp add: norm_mult) also have "\ = \r\" by auto finally have "cmod (z-z0) = \r\" . then show False using assms by auto qed lemma circlepath_inj_on: assumes "r>0" shows "inj_on (circlepath z r) {0..<1}" proof (rule inj_onI) fix x y assume asm: "x \ {0..<1}" "y \ {0..<1}" "circlepath z r x = circlepath z r y" define c where "c=2 * pi * \" have "c\0" unfolding c_def by auto from asm(3) have "exp (c * x) =exp (c * y)" unfolding circlepath c_def using \r>0\ by auto then obtain n where "c * x =c * (y + of_int n)" by (auto simp add:exp_eq c_def algebra_simps) then have "x=y+n" using \c\0\ by (meson mult_cancel_left of_real_eq_iff) then show "x=y" using asm(1,2) by auto qed subsection \More lemmas related to @{term winding_number}\ lemma winding_number_comp: assumes "open s" "f holomorphic_on s" "path_image \ \ s" "valid_path \" "z \ path_image (f \ \)" shows "winding_number (f \ \) z = 1/(2*pi*\)* contour_integral \ (\w. deriv f w / (f w - z))" proof - obtain spikes where "finite spikes" and \_diff: "\ C1_differentiable_on {0..1} - spikes" using \valid_path \\ unfolding valid_path_def piecewise_C1_differentiable_on_def by auto have "valid_path (f \ \)" using valid_path_compose_holomorphic assms by blast moreover have "contour_integral (f \ \) (\w. 1 / (w - z)) = contour_integral \ (\w. deriv f w / (f w - z))" unfolding contour_integral_integral proof (rule integral_spike[rule_format,OF negligible_finite[OF \finite spikes\]]) fix t::real assume t:"t \ {0..1} - spikes" then have "\ differentiable at t" using \_diff unfolding C1_differentiable_on_eq by auto moreover have "f field_differentiable at (\ t)" proof - have "\ t \ s" using \path_image \ \ s\ t unfolding path_image_def by auto thus ?thesis using \open s\ \f holomorphic_on s\ holomorphic_on_imp_differentiable_at by blast qed ultimately show " deriv f (\ t) / (f (\ t) - z) * vector_derivative \ (at t) = 1 / ((f \ \) t - z) * vector_derivative (f \ \) (at t)" apply (subst vector_derivative_chain_at_general) by (simp_all add:field_simps) qed moreover note \z \ path_image (f \ \)\ ultimately show ?thesis apply (subst winding_number_valid_path) by simp_all qed lemma winding_number_uminus_comp: assumes "valid_path \" "- z \ path_image \" shows "winding_number (uminus \ \) z = winding_number \ (-z)" proof - define c where "c= 2 * pi * \" have "winding_number (uminus \ \) z = 1/c * contour_integral \ (\w. deriv uminus w / (-w-z)) " proof (rule winding_number_comp[of UNIV, folded c_def]) show "open UNIV" "uminus holomorphic_on UNIV" "path_image \ \ UNIV" "valid_path \" using \valid_path \\ by (auto intro:holomorphic_intros) show "z \ path_image (uminus \ \)" unfolding path_image_compose using \- z \ path_image \\ by auto qed also have "\ = 1/c * contour_integral \ (\w. 1 / (w- (-z)))" - apply (auto intro!:contour_integral_eq simp add:field_simps) - apply (subst deriv_linear[of "-1",simplified]) - by (simp add: minus_divide_right) + by (auto intro!:contour_integral_eq simp add:field_simps minus_divide_right) also have "\ = winding_number \ (-z)" using winding_number_valid_path[OF \valid_path \\ \- z \ path_image \\,folded c_def] by simp finally show ?thesis by auto qed lemma winding_number_comp_linear: assumes "c\0" "valid_path \" and not_image: "(z-b)/c \ path_image \" shows "winding_number ((\x. c*x+b) \ \) z = winding_number \ ((z-b)/c)" (is "?L = ?R") proof - define cc where "cc=1 / (complex_of_real (2 * pi) * \)" define zz where "zz=(z-b)/c" have "?L = cc * contour_integral \ (\w. deriv (\x. c * x + b) w / (c * w + b - z))" apply (subst winding_number_comp[of UNIV,simplified]) subgoal by (auto intro:holomorphic_intros) subgoal using \valid_path \\ . subgoal using not_image \c\0\ unfolding path_image_compose by auto subgoal unfolding cc_def by auto done also have "\ = cc * contour_integral \ (\w.1 / (w - zz))" proof - have "deriv (\x. c * x + b) = (\x. c)" by (auto intro:derivative_intros) then show ?thesis unfolding zz_def cc_def using \c\0\ by (auto simp:field_simps) qed also have "\ = winding_number \ zz" using winding_number_valid_path[OF \valid_path \\ not_image,folded zz_def cc_def] by simp finally show "winding_number ((\x. c * x + b) \ \) z = winding_number \ zz" . qed end diff --git a/thys/Winding_Number_Eval/ROOT b/thys/Winding_Number_Eval/ROOT --- a/thys/Winding_Number_Eval/ROOT +++ b/thys/Winding_Number_Eval/ROOT @@ -1,15 +1,16 @@ chapter AFP -session "Winding_Number_Eval" (AFP) = "HOL-Analysis" + +session "Winding_Number_Eval" (AFP) = "HOL-Complex_Analysis" + options [timeout = 600] sessions + "HOL-Analysis" "HOL-Computational_Algebra" "HOL-Eisbach" "Sturm_Tarski" "Budan_Fourier" theories Winding_Number_Eval Winding_Number_Eval_Examples document_files "root.tex" "root.bib"